diff --git a/src/HOL/Fun.thy b/src/HOL/Fun.thy --- a/src/HOL/Fun.thy +++ b/src/HOL/Fun.thy @@ -1,1076 +1,1108 @@ (* 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 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 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) 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 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_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 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 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) 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 \\swap\\ definition swap :: "'a \ 'a \ ('a \ 'b) \ ('a \ 'b)" where "swap a b f = f (a := f b, b:= f a)" lemma swap_apply [simp]: "swap a b f a = f b" "swap a b f b = f a" "c \ a \ c \ b \ swap a b f c = f c" by (simp_all add: swap_def) lemma swap_self [simp]: "swap a a f = f" by (simp add: swap_def) lemma swap_commute: "swap a b f = swap b a f" by (simp add: fun_upd_def swap_def fun_eq_iff) lemma swap_nilpotent [simp]: "swap a b (swap a b f) = f" by (rule ext) (simp add: fun_upd_def swap_def) lemma swap_comp_involutory [simp]: "swap a b \ swap a b = id" by (rule ext) simp lemma swap_triple: assumes "a \ c" and "b \ c" shows "swap a b (swap b c (swap a b f)) = swap a c f" using assms by (simp add: fun_eq_iff swap_def) lemma comp_swap: "f \ swap a b g = swap a b (f \ g)" by (rule ext) (simp add: fun_upd_def swap_def) lemma swap_image_eq [simp]: assumes "a \ A" "b \ A" shows "swap a b f ` A = f ` A" proof - have subset: "\f. swap a b f ` A \ f ` A" using assms by (auto simp: image_iff swap_def) then have "swap a b (swap a b f) ` A \ (swap a b f) ` A" . with subset[of f] show ?thesis by auto qed lemma inj_on_imp_inj_on_swap: "inj_on f A \ a \ A \ b \ A \ inj_on (swap a b f) A" by (auto simp add: inj_on_def swap_def) lemma inj_on_swap_iff [simp]: assumes A: "a \ A" "b \ A" shows "inj_on (swap a b f) A \ inj_on f A" proof assume "inj_on (swap a b f) A" with A have "inj_on (swap a b (swap a b f)) A" by (iprover intro: inj_on_imp_inj_on_swap) then show "inj_on f A" by simp next assume "inj_on f A" with A show "inj_on (swap a b f) A" by (iprover intro: inj_on_imp_inj_on_swap) qed lemma surj_imp_surj_swap: "surj f \ surj (swap a b f)" by simp lemma surj_swap_iff [simp]: "surj (swap a b f) \ surj f" by simp lemma bij_betw_swap_iff [simp]: "x \ A \ y \ A \ bij_betw (swap x y f) A B \ bij_betw f A B" by (auto simp: bij_betw_def) lemma bij_swap_iff [simp]: "bij (swap a b f) \ bij f" by simp hide_const (open) swap 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) 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 "mono_on f A \ \r s. r \ A \ s \ A \ r \ s \ f r \ f s" lemma mono_onI: "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ mono_on f A" unfolding mono_on_def by simp lemma mono_onD: "\mono_on f A; r \ A; s \ A; r \ s\ \ f r \ f s" unfolding mono_on_def by simp lemma mono_imp_mono_on: "mono f \ mono_on f A" unfolding mono_def mono_on_def by auto lemma mono_on_subset: "mono_on f A \ B \ A \ mono_on f B" unfolding mono_on_def by auto definition "strict_mono_on f A \ \r s. r \ A \ s \ A \ r < s \ f r < f s" lemma strict_mono_onI: "(\r s. r \ A \ s \ A \ r < s \ f r < f s) \ strict_mono_on f A" unfolding strict_mono_on_def by simp lemma strict_mono_onD: "\strict_mono_on f A; r \ A; s \ A; r < s\ \ f r < f s" unfolding strict_mono_on_def by simp lemma mono_on_greaterD: assumes "mono_on g A" "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 (f :: (_ :: linorder) \ (_ :: preorder)) A" 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 (f :: (_ :: linorder) \ _ :: preorder) A" "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) lemma strict_mono_on_eqD: fixes f :: "(_ :: linorder) \ (_ :: preorder)" assumes "strict_mono_on f A" "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) lemma strict_mono_on_imp_mono_on: "strict_mono_on (f :: (_ :: linorder) \ _ :: preorder) A \ mono_on f A" 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/Lattices_Big.thy b/src/HOL/Lattices_Big.thy --- a/src/HOL/Lattices_Big.thy +++ b/src/HOL/Lattices_Big.thy @@ -1,1067 +1,1071 @@ (* 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 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 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 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 antisym) lemma Min_eq_iff: "\ finite A; A \ {} \ \ Min A = m \ m \ A \ (\a \ A. m \ a)" by (meson Min_in Min_le antisym) lemma eq_Max_iff: "\ finite A; A \ {} \ \ m = Max A \ m \ A \ (\a \ A. a \ m)" by (meson Max_in Max_ge antisym) lemma Max_eq_iff: "\ finite A; A \ {} \ \ Max A = m \ m \ A \ (\a \ A. a \ m)" by (meson Max_in Max_ge 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: 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 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 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)+ 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 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/Library/Disjoint_FSets.thy b/src/HOL/Library/Disjoint_FSets.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Library/Disjoint_FSets.thy @@ -0,0 +1,72 @@ +(* Title: HOL/Library/Disjoint_FSets.thy + Author: Lars Hupel, TU München +*) + +section \Disjoint FSets\ + +theory Disjoint_FSets + imports + "HOL-Library.Finite_Map" + "HOL-Library.Disjoint_Sets" +begin + +context + includes fset.lifting +begin + +lift_definition fdisjnt :: "'a fset \ 'a fset \ bool" is disjnt . + +lemma fdisjnt_alt_def: "fdisjnt M N \ (M |\| N = {||})" +by transfer (simp add: disjnt_def) + +lemma fdisjnt_insert: "x |\| N \ fdisjnt M N \ fdisjnt (finsert x M) N" +by transfer' (rule disjnt_insert) + +lemma fdisjnt_subset_right: "N' |\| N \ fdisjnt M N \ fdisjnt M N'" +unfolding fdisjnt_alt_def by auto + +lemma fdisjnt_subset_left: "N' |\| N \ fdisjnt N M \ fdisjnt N' M" +unfolding fdisjnt_alt_def by auto + +lemma fdisjnt_union_right: "fdisjnt M A \ fdisjnt M B \ fdisjnt M (A |\| B)" +unfolding fdisjnt_alt_def by auto + +lemma fdisjnt_union_left: "fdisjnt A M \ fdisjnt B M \ fdisjnt (A |\| B) M" +unfolding fdisjnt_alt_def by auto + +lemma fdisjnt_swap: "fdisjnt M N \ fdisjnt N M" +including fset.lifting by transfer' (auto simp: disjnt_def) + +lemma distinct_append_fset: + assumes "distinct xs" "distinct ys" "fdisjnt (fset_of_list xs) (fset_of_list ys)" + shows "distinct (xs @ ys)" +using assms +by transfer' (simp add: disjnt_def) + +lemma fdisjnt_contrI: + assumes "\x. x |\| M \ x |\| N \ False" + shows "fdisjnt M N" +using assms +by transfer' (auto simp: disjnt_def) + +lemma fdisjnt_Union_left: "fdisjnt (ffUnion S) T \ fBall S (\S. fdisjnt S T)" +by transfer' (auto simp: disjnt_def) + +lemma fdisjnt_Union_right: "fdisjnt T (ffUnion S) \ fBall S (\S. fdisjnt T S)" +by transfer' (auto simp: disjnt_def) + +lemma fdisjnt_ge_max: "fBall X (\x. x > fMax Y) \ fdisjnt X Y" +by transfer (auto intro: disjnt_ge_max) + +end + +(* FIXME should be provable without lifting *) +lemma fmadd_disjnt: "fdisjnt (fmdom m) (fmdom n) \ m ++\<^sub>f n = n ++\<^sub>f m" +unfolding fdisjnt_alt_def +including fset.lifting fmap.lifting +apply transfer +apply (rule ext) +apply (auto simp: map_add_def split: option.splits) +done + +end diff --git a/src/HOL/Library/Library.thy b/src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy +++ b/src/HOL/Library/Library.thy @@ -1,106 +1,107 @@ (*<*) theory Library imports AList Adhoc_Overloading BigO Bit_Operations BNF_Axiomatization BNF_Corec Boolean_Algebra Bourbaki_Witt_Fixpoint Char_ord Code_Lazy Code_Test Combine_PER Complete_Partial_Order2 Conditional_Parametricity Confluence Confluent_Quotient Countable Countable_Complete_Lattices Countable_Set_Type Debug Diagonal_Subsequence Discrete Disjoint_Sets + Disjoint_FSets Dlist Dual_Ordered_Lattice Equipollence Extended Extended_Nat Extended_Nonnegative_Real Extended_Real Finite_Map Float FSet FuncSet Function_Division Fun_Lexorder Going_To_Filter Groups_Big_Fun Indicator_Function Infinite_Set Interval Interval_Float IArray Landau_Symbols Lattice_Algebras Lattice_Syntax Lattice_Constructions Linear_Temporal_Logic_on_Streams ListVector List_Permutation Lub_Glb Mapping Monad_Syntax More_List Multiset_Order Multiset_Permutations Nonpos_Ints Numeral_Type Omega_Words_Fun Open_State_Syntax Option_ord Order_Continuity Parallel Pattern_Aliases Periodic_Fun Perm Permutations Poly_Mapping Power_By_Squaring Preorder Product_Plus Quadratic_Discriminant Quotient_List Quotient_Option Quotient_Product Quotient_Set Quotient_Sum Quotient_Syntax Quotient_Type Ramsey Reflection Rewrite Saturated Set_Algebras Set_Idioms Signed_Division State_Monad Stirling Stream Sorting_Algorithms Sublist Sum_of_Squares Transitive_Closure_Table Tree_Multiset Tree_Real Type_Length Uprod While_Combinator Word Z2 begin end (*>*) diff --git a/src/HOL/Set.thy b/src/HOL/Set.thy --- a/src/HOL/Set.thy +++ b/src/HOL/Set.thy @@ -1,2034 +1,2038 @@ (* Title: HOL/Set.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Set theory for higher-order logic\ theory Set imports Lattices begin subsection \Sets as predicates\ typedecl 'a set axiomatization Collect :: "('a \ bool) \ 'a set" \ \comprehension\ and member :: "'a \ 'a set \ bool" \ \membership\ where mem_Collect_eq [iff, code_unfold]: "member a (Collect P) = P a" and Collect_mem_eq [simp]: "Collect (\x. member x A) = A" notation member ("'(\')") and member ("(_/ \ _)" [51, 51] 50) abbreviation not_member where "not_member x A \ \ (x \ A)" \ \non-membership\ notation not_member ("'(\')") and not_member ("(_/ \ _)" [51, 51] 50) notation (ASCII) member ("'(:')") and member ("(_/ : _)" [51, 51] 50) and not_member ("'(~:')") and not_member ("(_/ ~: _)" [51, 51] 50) text \Set comprehensions\ syntax "_Coll" :: "pttrn \ bool \ 'a set" ("(1{_./ _})") translations "{x. P}" \ "CONST Collect (\x. P)" syntax (ASCII) "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{(_/: _)./ _})") syntax "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{(_/ \ _)./ _})") translations "{p:A. P}" \ "CONST Collect (\p. p \ A \ P)" lemma CollectI: "P a \ a \ {x. P x}" by simp lemma CollectD: "a \ {x. P x} \ P a" by simp lemma Collect_cong: "(\x. P x = Q x) \ {x. P x} = {x. Q x}" by simp text \ Simproc for pulling \x = t\ in \{x. \ \ x = t \ \}\ to the front (and similarly for \t = x\): \ simproc_setup defined_Collect ("{x. P x \ Q x}") = \ fn _ => Quantifier1.rearrange_Collect (fn ctxt => resolve_tac ctxt @{thms Collect_cong} 1 THEN resolve_tac ctxt @{thms iffI} 1 THEN ALLGOALS (EVERY' [REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}, DEPTH_SOLVE_1 o (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms conjI})])) \ lemmas CollectE = CollectD [elim_format] lemma set_eqI: assumes "\x. x \ A \ x \ B" shows "A = B" proof - from assms have "{x. x \ A} = {x. x \ B}" by simp then show ?thesis by simp qed lemma set_eq_iff: "A = B \ (\x. x \ A \ x \ B)" by (auto intro:set_eqI) lemma Collect_eqI: assumes "\x. P x = Q x" shows "Collect P = Collect Q" using assms by (auto intro: set_eqI) text \Lifting of predicate class instances\ instantiation set :: (type) boolean_algebra begin definition less_eq_set where "A \ B \ (\x. member x A) \ (\x. member x B)" definition less_set where "A < B \ (\x. member x A) < (\x. member x B)" definition inf_set where "A \ B = Collect ((\x. member x A) \ (\x. member x B))" definition sup_set where "A \ B = Collect ((\x. member x A) \ (\x. member x B))" definition bot_set where "\ = Collect \" definition top_set where "\ = Collect \" definition uminus_set where "- A = Collect (- (\x. member x A))" definition minus_set where "A - B = Collect ((\x. member x A) - (\x. member x B))" instance by standard (simp_all add: less_eq_set_def less_set_def inf_set_def sup_set_def bot_set_def top_set_def uminus_set_def minus_set_def less_le_not_le sup_inf_distrib1 diff_eq set_eqI fun_eq_iff del: inf_apply sup_apply bot_apply top_apply minus_apply uminus_apply) end text \Set enumerations\ abbreviation empty :: "'a set" ("{}") where "{} \ bot" definition insert :: "'a \ 'a set \ 'a set" where insert_compr: "insert a B = {x. x = a \ x \ B}" syntax "_Finset" :: "args \ 'a set" ("{(_)}") translations "{x, xs}" \ "CONST insert x {xs}" "{x}" \ "CONST insert x {}" subsection \Subsets and bounded quantifiers\ abbreviation subset :: "'a set \ 'a set \ bool" where "subset \ less" abbreviation subset_eq :: "'a set \ 'a set \ bool" where "subset_eq \ less_eq" notation subset ("'(\')") and subset ("(_/ \ _)" [51, 51] 50) and subset_eq ("'(\')") and subset_eq ("(_/ \ _)" [51, 51] 50) abbreviation (input) supset :: "'a set \ 'a set \ bool" where "supset \ greater" abbreviation (input) supset_eq :: "'a set \ 'a set \ bool" where "supset_eq \ greater_eq" notation supset ("'(\')") and supset ("(_/ \ _)" [51, 51] 50) and supset_eq ("'(\')") and supset_eq ("(_/ \ _)" [51, 51] 50) notation (ASCII output) subset ("'(<')") and subset ("(_/ < _)" [51, 51] 50) and subset_eq ("'(<=')") and subset_eq ("(_/ <= _)" [51, 51] 50) definition Ball :: "'a set \ ('a \ bool) \ bool" where "Ball A P \ (\x. x \ A \ P x)" \ \bounded universal quantifiers\ definition Bex :: "'a set \ ('a \ bool) \ bool" where "Bex A P \ (\x. x \ A \ P x)" \ \bounded existential quantifiers\ syntax (ASCII) "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3ALL (_/:_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3EX (_/:_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3EX! (_/:_)./ _)" [0, 0, 10] 10) "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST (_/:_)./ _)" [0, 0, 10] 10) syntax (input) "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3! (_/:_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3? (_/:_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3?! (_/:_)./ _)" [0, 0, 10] 10) syntax "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3\(_/\_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3\(_/\_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3\!(_/\_)./ _)" [0, 0, 10] 10) "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST(_/\_)./ _)" [0, 0, 10] 10) translations "\x\A. P" \ "CONST Ball A (\x. P)" "\x\A. P" \ "CONST Bex A (\x. P)" "\!x\A. P" \ "\!x. x \ A \ P" "LEAST x:A. P" \ "LEAST x. x \ A \ P" syntax (ASCII output) "_setlessAll" :: "[idt, 'a, bool] \ bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) "_setlessEx" :: "[idt, 'a, bool] \ bool" ("(3EX _<_./ _)" [0, 0, 10] 10) "_setleAll" :: "[idt, 'a, bool] \ bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3EX! _<=_./ _)" [0, 0, 10] 10) syntax "_setlessAll" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setlessEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleAll" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3\!_\_./ _)" [0, 0, 10] 10) translations "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\!A\B. P" \ "\!A. A \ B \ 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 sbset = \<^const_syntax>\subset\; val sbset_eq = \<^const_syntax>\subset_eq\; val trans = [((All_binder, impl, sbset), \<^syntax_const>\_setlessAll\), ((All_binder, impl, sbset_eq), \<^syntax_const>\_setleAll\), ((Ex_binder, conj, sbset), \<^syntax_const>\_setlessEx\), ((Ex_binder, conj, sbset_eq), \<^syntax_const>\_setleEx\)]; fun mk v (v', T) c n P = if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n) then Syntax.const c $ Syntax_Trans.mark_bound_body (v', T) $ n $ P else raise Match; fun tr' q = (q, fn _ => (fn [Const (\<^syntax_const>\_bound\, _) $ Free (v, Type (\<^type_name>\set\, _)), Const (c, _) $ (Const (d, _) $ (Const (\<^syntax_const>\_bound\, _) $ Free (v', T)) $ n) $ P] => (case AList.lookup (=) trans (q, c, d) of NONE => raise Match | SOME l => mk v (v', T) l n P) | _ => raise Match)); in [tr' All_binder, tr' Ex_binder] end \ text \ \<^medskip> Translate between \{e | x1\xn. P}\ and \{u. \x1\xn. u = e \ P}\; \{y. \x1\xn. y = e \ P}\ is only translated if \[0..n] \ bvs e\. \ syntax "_Setcompr" :: "'a \ idts \ bool \ 'a set" ("(1{_ |/_./ _})") parse_translation \ let val ex_tr = snd (Syntax_Trans.mk_binder_tr ("EX ", \<^const_syntax>\Ex\)); fun nvars (Const (\<^syntax_const>\_idts\, _) $ _ $ idts) = nvars idts + 1 | nvars _ = 1; fun setcompr_tr ctxt [e, idts, b] = let val eq = Syntax.const \<^const_syntax>\HOL.eq\ $ Bound (nvars idts) $ e; val P = Syntax.const \<^const_syntax>\HOL.conj\ $ eq $ b; val exP = ex_tr ctxt [idts, P]; in Syntax.const \<^const_syntax>\Collect\ $ absdummy dummyT exP end; in [(\<^syntax_const>\_Setcompr\, setcompr_tr)] end \ print_translation \ [Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Ball\ \<^syntax_const>\_Ball\, Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Bex\ \<^syntax_const>\_Bex\] \ \ \to avoid eta-contraction of body\ print_translation \ let val ex_tr' = snd (Syntax_Trans.mk_binder_tr' (\<^const_syntax>\Ex\, "DUMMY")); fun setcompr_tr' ctxt [Abs (abs as (_, _, P))] = let fun check (Const (\<^const_syntax>\Ex\, _) $ Abs (_, _, P), n) = check (P, n + 1) | check (Const (\<^const_syntax>\HOL.conj\, _) $ (Const (\<^const_syntax>\HOL.eq\, _) $ Bound m $ e) $ P, n) = n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso subset (=) (0 upto (n - 1), add_loose_bnos (e, 0, [])) | check _ = false; fun tr' (_ $ abs) = let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' ctxt [abs] in Syntax.const \<^syntax_const>\_Setcompr\ $ e $ idts $ Q end; in if check (P, 0) then tr' P else let val (x as _ $ Free(xN, _), t) = Syntax_Trans.atomic_abs_tr' abs; val M = Syntax.const \<^syntax_const>\_Coll\ $ x $ t; in case t of Const (\<^const_syntax>\HOL.conj\, _) $ (Const (\<^const_syntax>\Set.member\, _) $ (Const (\<^syntax_const>\_bound\, _) $ Free (yN, _)) $ A) $ P => if xN = yN then Syntax.const \<^syntax_const>\_Collect\ $ x $ A $ P else M | _ => M end end; in [(\<^const_syntax>\Collect\, setcompr_tr')] end \ simproc_setup defined_Bex ("\x\A. P x \ Q x") = \ fn _ => Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def}) \ simproc_setup defined_All ("\x\A. P x \ Q x") = \ fn _ => Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def}) \ lemma ballI [intro!]: "(\x. x \ A \ P x) \ \x\A. P x" by (simp add: Ball_def) lemmas strip = impI allI ballI lemma bspec [dest?]: "\x\A. P x \ x \ A \ P x" by (simp add: Ball_def) text \Gives better instantiation for bound:\ setup \ map_theory_claset (fn ctxt => ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt')) \ ML \ structure Simpdata = struct open Simpdata; val mksimps_pairs = [(\<^const_name>\Ball\, @{thms bspec})] @ mksimps_pairs; end; open Simpdata; \ declaration \fn _ => Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))\ lemma ballE [elim]: "\x\A. P x \ (P x \ Q) \ (x \ A \ Q) \ Q" unfolding Ball_def by blast lemma bexI [intro]: "P x \ x \ A \ \x\A. P x" \ \Normally the best argument order: \P x\ constrains the choice of \x \ A\.\ unfolding Bex_def by blast lemma rev_bexI [intro?]: "x \ A \ P x \ \x\A. P x" \ \The best argument order when there is only one \x \ A\.\ unfolding Bex_def by blast lemma bexCI: "(\x\A. \ P x \ P a) \ a \ A \ \x\A. P x" unfolding Bex_def by blast lemma bexE [elim!]: "\x\A. P x \ (\x. x \ A \ P x \ Q) \ Q" unfolding Bex_def by blast lemma ball_triv [simp]: "(\x\A. P) \ ((\x. x \ A) \ P)" \ \trivial rewrite rule.\ by (simp add: Ball_def) lemma bex_triv [simp]: "(\x\A. P) \ ((\x. x \ A) \ P)" \ \Dual form for existentials.\ by (simp add: Bex_def) lemma bex_triv_one_point1 [simp]: "(\x\A. x = a) \ a \ A" by blast lemma bex_triv_one_point2 [simp]: "(\x\A. a = x) \ a \ A" by blast lemma bex_one_point1 [simp]: "(\x\A. x = a \ P x) \ a \ A \ P a" by blast lemma bex_one_point2 [simp]: "(\x\A. a = x \ P x) \ a \ A \ P a" by blast lemma ball_one_point1 [simp]: "(\x\A. x = a \ P x) \ (a \ A \ P a)" by blast lemma ball_one_point2 [simp]: "(\x\A. a = x \ P x) \ (a \ A \ P a)" by blast lemma ball_conj_distrib: "(\x\A. P x \ Q x) \ (\x\A. P x) \ (\x\A. Q x)" by blast lemma bex_disj_distrib: "(\x\A. P x \ Q x) \ (\x\A. P x) \ (\x\A. Q x)" by blast text \Congruence rules\ lemma ball_cong: "\ A = B; \x. x \ B \ P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: Ball_def) lemma ball_cong_simp [cong]: "\ A = B; \x. x \ B =simp=> P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Ball_def) lemma bex_cong: "\ A = B; \x. x \ B \ P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: Bex_def cong: conj_cong) lemma bex_cong_simp [cong]: "\ A = B; \x. x \ B =simp=> P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Bex_def cong: conj_cong) lemma bex1_def: "(\!x\X. P x) \ (\x\X. P x) \ (\x\X. \y\X. P x \ P y \ x = y)" by auto subsection \Basic operations\ subsubsection \Subsets\ lemma subsetI [intro!]: "(\x. x \ A \ x \ B) \ A \ B" by (simp add: less_eq_set_def le_fun_def) text \ \<^medskip> Map the type \'a set \ anything\ to just \'a\; for overloading constants whose first argument has type \'a set\. \ lemma subsetD [elim, intro?]: "A \ B \ c \ A \ c \ B" by (simp add: less_eq_set_def le_fun_def) \ \Rule in Modus Ponens style.\ lemma rev_subsetD [intro?,no_atp]: "c \ A \ A \ B \ c \ B" \ \The same, with reversed premises for use with @{method erule} -- cf. @{thm rev_mp}.\ by (rule subsetD) lemma subsetCE [elim,no_atp]: "A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" \ \Classical elimination rule.\ by (auto simp add: less_eq_set_def le_fun_def) lemma subset_eq: "A \ B \ (\x\A. x \ B)" by blast lemma contra_subsetD [no_atp]: "A \ B \ c \ B \ c \ A" by blast lemma subset_refl: "A \ A" by (fact order_refl) (* already [iff] *) lemma subset_trans: "A \ B \ B \ C \ A \ C" by (fact order_trans) lemma subset_not_subset_eq [code]: "A \ B \ A \ B \ \ B \ A" by (fact less_le_not_le) lemma eq_mem_trans: "a = b \ b \ A \ a \ A" by simp lemmas basic_trans_rules [trans] = order_trans_rules rev_subsetD subsetD eq_mem_trans subsubsection \Equality\ lemma subset_antisym [intro!]: "A \ B \ B \ A \ A = B" \ \Anti-symmetry of the subset relation.\ by (iprover intro: set_eqI subsetD) text \\<^medskip> Equality rules from ZF set theory -- are they appropriate here?\ lemma equalityD1: "A = B \ A \ B" by simp lemma equalityD2: "A = B \ B \ A" by simp text \ \<^medskip> Be careful when adding this to the claset as \subset_empty\ is in the simpset: \<^prop>\A = {}\ goes to \<^prop>\{} \ A\ and \<^prop>\A \ {}\ and then back to \<^prop>\A = {}\! \ lemma equalityE: "A = B \ (A \ B \ B \ A \ P) \ P" by simp lemma equalityCE [elim]: "A = B \ (c \ A \ c \ B \ P) \ (c \ A \ c \ B \ P) \ P" by blast lemma eqset_imp_iff: "A = B \ x \ A \ x \ B" by simp lemma eqelem_imp_iff: "x = y \ x \ A \ y \ A" by simp subsubsection \The empty set\ lemma empty_def: "{} = {x. False}" by (simp add: bot_set_def bot_fun_def) lemma empty_iff [simp]: "c \ {} \ False" by (simp add: empty_def) lemma emptyE [elim!]: "a \ {} \ P" by simp lemma empty_subsetI [iff]: "{} \ A" \ \One effect is to delete the ASSUMPTION \<^prop>\{} \ A\\ by blast lemma equals0I: "(\y. y \ A \ False) \ A = {}" by blast lemma equals0D: "A = {} \ a \ A" \ \Use for reasoning about disjointness: \A \ B = {}\\ by blast lemma ball_empty [simp]: "Ball {} P \ True" by (simp add: Ball_def) lemma bex_empty [simp]: "Bex {} P \ False" by (simp add: Bex_def) subsubsection \The universal set -- UNIV\ abbreviation UNIV :: "'a set" where "UNIV \ top" lemma UNIV_def: "UNIV = {x. True}" by (simp add: top_set_def top_fun_def) lemma UNIV_I [simp]: "x \ UNIV" by (simp add: UNIV_def) declare UNIV_I [intro] \ \unsafe makes it less likely to cause problems\ lemma UNIV_witness [intro?]: "\x. x \ UNIV" by simp lemma subset_UNIV: "A \ UNIV" by (fact top_greatest) (* already simp *) text \ \<^medskip> Eta-contracting these two rules (to remove \P\) causes them to be ignored because of their interaction with congruence rules. \ lemma ball_UNIV [simp]: "Ball UNIV P \ All P" by (simp add: Ball_def) lemma bex_UNIV [simp]: "Bex UNIV P \ Ex P" by (simp add: Bex_def) lemma UNIV_eq_I: "(\x. x \ A) \ UNIV = A" by auto lemma UNIV_not_empty [iff]: "UNIV \ {}" by (blast elim: equalityE) lemma empty_not_UNIV[simp]: "{} \ UNIV" by blast subsubsection \The Powerset operator -- Pow\ definition Pow :: "'a set \ 'a set set" where Pow_def: "Pow A = {B. B \ A}" lemma Pow_iff [iff]: "A \ Pow B \ A \ B" by (simp add: Pow_def) lemma PowI: "A \ B \ A \ Pow B" by (simp add: Pow_def) lemma PowD: "A \ Pow B \ A \ B" by (simp add: Pow_def) lemma Pow_bottom: "{} \ Pow B" by simp lemma Pow_top: "A \ Pow A" by simp lemma Pow_not_empty: "Pow A \ {}" using Pow_top by blast subsubsection \Set complement\ lemma Compl_iff [simp]: "c \ - A \ c \ A" by (simp add: fun_Compl_def uminus_set_def) lemma ComplI [intro!]: "(c \ A \ False) \ c \ - A" by (simp add: fun_Compl_def uminus_set_def) blast text \ \<^medskip> This form, with negated conclusion, works well with the Classical prover. Negated assumptions behave like formulae on the right side of the notional turnstile \dots \ lemma ComplD [dest!]: "c \ - A \ c \ A" by simp lemmas ComplE = ComplD [elim_format] lemma Compl_eq: "- A = {x. \ x \ A}" by blast subsubsection \Binary intersection\ abbreviation inter :: "'a set \ 'a set \ 'a set" (infixl "\" 70) where "(\) \ inf" notation (ASCII) inter (infixl "Int" 70) lemma Int_def: "A \ B = {x. x \ A \ x \ B}" by (simp add: inf_set_def inf_fun_def) lemma Int_iff [simp]: "c \ A \ B \ c \ A \ c \ B" unfolding Int_def by blast lemma IntI [intro!]: "c \ A \ c \ B \ c \ A \ B" by simp lemma IntD1: "c \ A \ B \ c \ A" by simp lemma IntD2: "c \ A \ B \ c \ B" by simp lemma IntE [elim!]: "c \ A \ B \ (c \ A \ c \ B \ P) \ P" by simp lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" by (fact mono_inf) subsubsection \Binary union\ abbreviation union :: "'a set \ 'a set \ 'a set" (infixl "\" 65) where "union \ sup" notation (ASCII) union (infixl "Un" 65) lemma Un_def: "A \ B = {x. x \ A \ x \ B}" by (simp add: sup_set_def sup_fun_def) lemma Un_iff [simp]: "c \ A \ B \ c \ A \ c \ B" unfolding Un_def by blast lemma UnI1 [elim?]: "c \ A \ c \ A \ B" by simp lemma UnI2 [elim?]: "c \ B \ c \ A \ B" by simp text \\<^medskip> Classical introduction rule: no commitment to \A\ vs. \B\.\ lemma UnCI [intro!]: "(c \ B \ c \ A) \ c \ A \ B" by auto lemma UnE [elim!]: "c \ A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" unfolding Un_def by blast lemma insert_def: "insert a B = {x. x = a} \ B" by (simp add: insert_compr Un_def) lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" by (fact mono_sup) subsubsection \Set difference\ lemma Diff_iff [simp]: "c \ A - B \ c \ A \ c \ B" by (simp add: minus_set_def fun_diff_def) lemma DiffI [intro!]: "c \ A \ c \ B \ c \ A - B" by simp lemma DiffD1: "c \ A - B \ c \ A" by simp lemma DiffD2: "c \ A - B \ c \ B \ P" by simp lemma DiffE [elim!]: "c \ A - B \ (c \ A \ c \ B \ P) \ P" by simp lemma set_diff_eq: "A - B = {x. x \ A \ x \ B}" by blast lemma Compl_eq_Diff_UNIV: "- A = (UNIV - A)" by blast subsubsection \Augmenting a set -- \<^const>\insert\\ lemma insert_iff [simp]: "a \ insert b A \ a = b \ a \ A" unfolding insert_def by blast lemma insertI1: "a \ insert a B" by simp lemma insertI2: "a \ B \ a \ insert b B" by simp lemma insertE [elim!]: "a \ insert b A \ (a = b \ P) \ (a \ A \ P) \ P" unfolding insert_def by blast lemma insertCI [intro!]: "(a \ B \ a = b) \ a \ insert b B" \ \Classical introduction rule.\ by auto lemma subset_insert_iff: "A \ insert x B \ (if x \ A then A - {x} \ B else A \ B)" by auto lemma set_insert: assumes "x \ A" obtains B where "A = insert x B" and "x \ B" proof show "A = insert x (A - {x})" using assms by blast show "x \ A - {x}" by blast qed lemma insert_ident: "x \ A \ x \ B \ insert x A = insert x B \ A = B" by auto lemma insert_eq_iff: assumes "a \ A" "b \ B" shows "insert a A = insert b B \ (if a = b then A = B else \C. A = insert b C \ b \ C \ B = insert a C \ a \ C)" (is "?L \ ?R") proof show ?R if ?L proof (cases "a = b") case True with assms \?L\ show ?R by (simp add: insert_ident) next case False let ?C = "A - {b}" have "A = insert b ?C \ b \ ?C \ B = insert a ?C \ a \ ?C" using assms \?L\ \a \ b\ by auto then show ?R using \a \ b\ by auto qed show ?L if ?R using that by (auto split: if_splits) qed lemma insert_UNIV: "insert x UNIV = UNIV" by auto subsubsection \Singletons, using insert\ lemma singletonI [intro!]: "a \ {a}" \ \Redundant? But unlike \insertCI\, it proves the subgoal immediately!\ by (rule insertI1) lemma singletonD [dest!]: "b \ {a} \ b = a" by blast lemmas singletonE = singletonD [elim_format] lemma singleton_iff: "b \ {a} \ b = a" by blast lemma singleton_inject [dest!]: "{a} = {b} \ a = b" by blast lemma singleton_insert_inj_eq [iff]: "{b} = insert a A \ a = b \ A \ {b}" by blast lemma singleton_insert_inj_eq' [iff]: "insert a A = {b} \ a = b \ A \ {b}" by blast lemma subset_singletonD: "A \ {x} \ A = {} \ A = {x}" by fast lemma subset_singleton_iff: "X \ {a} \ X = {} \ X = {a}" by blast lemma subset_singleton_iff_Uniq: "(\a. A \ {a}) \ (\\<^sub>\\<^sub>1x. x \ A)" unfolding Uniq_def by blast lemma singleton_conv [simp]: "{x. x = a} = {a}" by blast lemma singleton_conv2 [simp]: "{x. a = x} = {a}" by blast lemma Diff_single_insert: "A - {x} \ B \ A \ insert x B" by blast lemma subset_Diff_insert: "A \ B - insert x C \ A \ B - C \ x \ A" by blast lemma doubleton_eq_iff: "{a, b} = {c, d} \ a = c \ b = d \ a = d \ b = c" by (blast elim: equalityE) lemma Un_singleton_iff: "A \ B = {x} \ A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x}" by auto lemma singleton_Un_iff: "{x} = A \ B \ A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x}" by auto subsubsection \Image of a set under a function\ text \Frequently \b\ does not have the syntactic form of \f x\.\ definition image :: "('a \ 'b) \ 'a set \ 'b set" (infixr "`" 90) where "f ` A = {y. \x\A. y = f x}" lemma image_eqI [simp, intro]: "b = f x \ x \ A \ b \ f ` A" unfolding image_def by blast lemma imageI: "x \ A \ f x \ f ` A" by (rule image_eqI) (rule refl) lemma rev_image_eqI: "x \ A \ b = f x \ b \ f ` A" \ \This version's more effective when we already have the required \x\.\ by (rule image_eqI) lemma imageE [elim!]: assumes "b \ (\x. f x) ` A" \ \The eta-expansion gives variable-name preservation.\ obtains x where "b = f x" and "x \ A" using assms unfolding image_def by blast lemma Compr_image_eq: "{x \ f ` A. P x} = f ` {x \ A. P (f x)}" by auto lemma image_Un: "f ` (A \ B) = f ` A \ f ` B" by blast lemma image_iff: "z \ f ` A \ (\x\A. z = f x)" by blast lemma image_subsetI: "(\x. x \ A \ f x \ B) \ f ` A \ B" \ \Replaces the three steps \subsetI\, \imageE\, \hypsubst\, but breaks too many existing proofs.\ by blast lemma image_subset_iff: "f ` A \ B \ (\x\A. f x \ B)" \ \This rewrite rule would confuse users if made default.\ by blast lemma subset_imageE: assumes "B \ f ` A" obtains C where "C \ A" and "B = f ` C" proof - from assms have "B = f ` {a \ A. f a \ B}" by fast moreover have "{a \ A. f a \ B} \ A" by blast ultimately show thesis by (blast intro: that) qed lemma subset_image_iff: "B \ f ` A \ (\AA\A. B = f ` AA)" by (blast elim: subset_imageE) lemma image_ident [simp]: "(\x. x) ` Y = Y" by blast lemma image_empty [simp]: "f ` {} = {}" by blast lemma image_insert [simp]: "f ` insert a B = insert (f a) (f ` B)" by blast lemma image_constant: "x \ A \ (\x. c) ` A = {c}" by auto lemma image_constant_conv: "(\x. c) ` A = (if A = {} then {} else {c})" by auto lemma image_image: "f ` (g ` A) = (\x. f (g x)) ` A" by blast lemma insert_image [simp]: "x \ A \ insert (f x) (f ` A) = f ` A" by blast lemma image_is_empty [iff]: "f ` A = {} \ A = {}" by blast lemma empty_is_image [iff]: "{} = f ` A \ A = {}" by blast lemma image_Collect: "f ` {x. P x} = {f x | x. P x}" \ \NOT suitable as a default simp rule: the RHS isn't simpler than the LHS, with its implicit quantifier and conjunction. Also image enjoys better equational properties than does the RHS.\ by blast lemma if_image_distrib [simp]: "(\x. if P x then f x else g x) ` S = f ` (S \ {x. P x}) \ g ` (S \ {x. \ P x})" by auto lemma image_cong: "f ` M = g ` N" if "M = N" "\x. x \ N \ f x = g x" using that by (simp add: image_def) lemma image_cong_simp [cong]: "f ` M = g ` N" if "M = N" "\x. x \ N =simp=> f x = g x" using that image_cong [of M N f g] by (simp add: simp_implies_def) lemma image_Int_subset: "f ` (A \ B) \ f ` A \ f ` B" by blast lemma image_diff_subset: "f ` A - f ` B \ f ` (A - B)" by blast lemma Setcompr_eq_image: "{f x |x. x \ A} = f ` A" by blast lemma setcompr_eq_image: "{f x |x. P x} = f ` {x. P x}" by auto lemma ball_imageD: "\x\f ` A. P x \ \x\A. P (f x)" by simp lemma bex_imageD: "\x\f ` A. P x \ \x\A. P (f x)" by auto lemma image_add_0 [simp]: "(+) (0::'a::comm_monoid_add) ` S = S" by auto text \\<^medskip> Range of a function -- just an abbreviation for image!\ abbreviation range :: "('a \ 'b) \ 'b set" \ \of function\ where "range f \ f ` UNIV" lemma range_eqI: "b = f x \ b \ range f" by simp lemma rangeI: "f x \ range f" by simp lemma rangeE [elim?]: "b \ range (\x. f x) \ (\x. b = f x \ P) \ P" by (rule imageE) lemma full_SetCompr_eq: "{u. \x. u = f x} = range f" by auto lemma range_composition: "range (\x. f (g x)) = f ` range g" by auto lemma range_constant [simp]: "range (\_. x) = {x}" by (simp add: image_constant) lemma range_eq_singletonD: "range f = {a} \ f x = a" by auto subsubsection \Some rules with \if\\ text \Elimination of \{x. \ \ x = t \ \}\.\ lemma Collect_conv_if: "{x. x = a \ P x} = (if P a then {a} else {})" by auto lemma Collect_conv_if2: "{x. a = x \ P x} = (if P a then {a} else {})" by auto text \ Rewrite rules for boolean case-splitting: faster than \if_split [split]\. \ lemma if_split_eq1: "(if Q then x else y) = b \ (Q \ x = b) \ (\ Q \ y = b)" by (rule if_split) lemma if_split_eq2: "a = (if Q then x else y) \ (Q \ a = x) \ (\ Q \ a = y)" by (rule if_split) text \ Split ifs on either side of the membership relation. Not for \[simp]\ -- can cause goals to blow up! \ lemma if_split_mem1: "(if Q then x else y) \ b \ (Q \ x \ b) \ (\ Q \ y \ b)" by (rule if_split) lemma if_split_mem2: "(a \ (if Q then x else y)) \ (Q \ a \ x) \ (\ Q \ a \ y)" by (rule if_split [where P = "\S. a \ S"]) lemmas split_ifs = if_bool_eq_conj if_split_eq1 if_split_eq2 if_split_mem1 if_split_mem2 (*Would like to add these, but the existing code only searches for the outer-level constant, which in this case is just Set.member; we instead need to use term-nets to associate patterns with rules. Also, if a rule fails to apply, then the formula should be kept. [("uminus", Compl_iff RS iffD1), ("minus", [Diff_iff RS iffD1]), ("Int", [IntD1,IntD2]), ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] *) subsection \Further operations and lemmas\ subsubsection \The ``proper subset'' relation\ lemma psubsetI [intro!]: "A \ B \ A \ B \ A \ B" unfolding less_le by blast lemma psubsetE [elim!]: "A \ B \ (A \ B \ \ B \ A \ R) \ R" unfolding less_le by blast lemma psubset_insert_iff: "A \ insert x B \ (if x \ B then A \ B else if x \ A then A - {x} \ B else A \ B)" by (auto simp add: less_le subset_insert_iff) lemma psubset_eq: "A \ B \ A \ B \ A \ B" by (simp only: less_le) lemma psubset_imp_subset: "A \ B \ A \ B" by (simp add: psubset_eq) lemma psubset_trans: "A \ B \ B \ C \ A \ C" unfolding less_le by (auto dest: subset_antisym) lemma psubsetD: "A \ B \ c \ A \ c \ B" unfolding less_le by (auto dest: subsetD) lemma psubset_subset_trans: "A \ B \ B \ C \ A \ C" by (auto simp add: psubset_eq) lemma subset_psubset_trans: "A \ B \ B \ C \ A \ C" by (auto simp add: psubset_eq) lemma psubset_imp_ex_mem: "A \ B \ \b. b \ B - A" unfolding less_le by blast lemma atomize_ball: "(\x. x \ A \ P x) \ Trueprop (\x\A. P x)" by (simp only: Ball_def atomize_all atomize_imp) lemmas [symmetric, rulify] = atomize_ball and [symmetric, defn] = atomize_ball lemma image_Pow_mono: "f ` A \ B \ image f ` Pow A \ Pow B" by blast lemma image_Pow_surj: "f ` A = B \ image f ` Pow A = Pow B" by (blast elim: subset_imageE) subsubsection \Derived rules involving subsets.\ text \\insert\.\ lemma subset_insertI: "B \ insert a B" by (rule subsetI) (erule insertI2) lemma subset_insertI2: "A \ B \ A \ insert b B" by blast lemma subset_insert: "x \ A \ A \ insert x B \ A \ B" by blast text \\<^medskip> Finite Union -- the least upper bound of two sets.\ lemma Un_upper1: "A \ A \ B" by (fact sup_ge1) lemma Un_upper2: "B \ A \ B" by (fact sup_ge2) lemma Un_least: "A \ C \ B \ C \ A \ B \ C" by (fact sup_least) text \\<^medskip> Finite Intersection -- the greatest lower bound of two sets.\ lemma Int_lower1: "A \ B \ A" by (fact inf_le1) lemma Int_lower2: "A \ B \ B" by (fact inf_le2) lemma Int_greatest: "C \ A \ C \ B \ C \ A \ B" by (fact inf_greatest) text \\<^medskip> Set difference.\ lemma Diff_subset[simp]: "A - B \ A" by blast lemma Diff_subset_conv: "A - B \ C \ A \ B \ C" by blast subsubsection \Equalities involving union, intersection, inclusion, etc.\ text \\{}\.\ lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})" \ \supersedes \Collect_False_empty\\ by auto lemma subset_empty [simp]: "A \ {} \ A = {}" by (fact bot_unique) lemma not_psubset_empty [iff]: "\ (A < {})" by (fact not_less_bot) (* FIXME: already simp *) lemma Collect_subset [simp]: "{x\A. P x} \ A" by auto lemma Collect_empty_eq [simp]: "Collect P = {} \ (\x. \ P x)" by blast lemma empty_Collect_eq [simp]: "{} = Collect P \ (\x. \ P x)" by blast lemma Collect_neg_eq: "{x. \ P x} = - {x. P x}" by blast lemma Collect_disj_eq: "{x. P x \ Q x} = {x. P x} \ {x. Q x}" by blast lemma Collect_imp_eq: "{x. P x \ Q x} = - {x. P x} \ {x. Q x}" by blast lemma Collect_conj_eq: "{x. P x \ Q x} = {x. P x} \ {x. Q x}" by blast lemma Collect_mono_iff: "Collect P \ Collect Q \ (\x. P x \ Q x)" by blast text \\<^medskip> \insert\.\ lemma insert_is_Un: "insert a A = {a} \ A" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a {}\\ by blast lemma insert_not_empty [simp]: "insert a A \ {}" and empty_not_insert [simp]: "{} \ insert a A" by blast+ lemma insert_absorb: "a \ A \ insert a A = A" \ \\[simp]\ causes recursive calls when there are nested inserts\ \ \with \<^emph>\quadratic\ running time\ by blast lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A" by blast lemma insert_commute: "insert x (insert y A) = insert y (insert x A)" by blast lemma insert_subset [simp]: "insert x A \ B \ x \ B \ A \ B" by blast lemma mk_disjoint_insert: "a \ A \ \B. A = insert a B \ a \ B" \ \use new \B\ rather than \A - {a}\ to avoid infinite unfolding\ by (rule exI [where x = "A - {a}"]) blast lemma insert_Collect: "insert a (Collect P) = {u. u \ a \ P u}" by auto lemma insert_inter_insert [simp]: "insert a A \ insert a B = insert a (A \ B)" by blast lemma insert_disjoint [simp]: "insert a A \ B = {} \ a \ B \ A \ B = {}" "{} = insert a A \ B \ a \ B \ {} = A \ B" by auto lemma disjoint_insert [simp]: "B \ insert a A = {} \ a \ B \ B \ A = {}" "{} = A \ insert b B \ b \ A \ {} = A \ B" by auto text \\<^medskip> \Int\\ lemma Int_absorb: "A \ A = A" by (fact inf_idem) (* already simp *) lemma Int_left_absorb: "A \ (A \ B) = A \ B" by (fact inf_left_idem) lemma Int_commute: "A \ B = B \ A" by (fact inf_commute) lemma Int_left_commute: "A \ (B \ C) = B \ (A \ C)" by (fact inf_left_commute) lemma Int_assoc: "(A \ B) \ C = A \ (B \ C)" by (fact inf_assoc) lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute \ \Intersection is an AC-operator\ lemma Int_absorb1: "B \ A \ A \ B = B" by (fact inf_absorb2) lemma Int_absorb2: "A \ B \ A \ B = A" by (fact inf_absorb1) lemma Int_empty_left: "{} \ B = {}" by (fact inf_bot_left) (* already simp *) lemma Int_empty_right: "A \ {} = {}" by (fact inf_bot_right) (* already simp *) lemma disjoint_eq_subset_Compl: "A \ B = {} \ A \ - B" by blast lemma disjoint_iff: "A \ B = {} \ (\x. x\A \ x \ B)" by blast lemma disjoint_iff_not_equal: "A \ B = {} \ (\x\A. \y\B. x \ y)" by blast lemma Int_UNIV_left: "UNIV \ B = B" by (fact inf_top_left) (* already simp *) lemma Int_UNIV_right: "A \ UNIV = A" by (fact inf_top_right) (* already simp *) lemma Int_Un_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by (fact inf_sup_distrib1) lemma Int_Un_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by (fact inf_sup_distrib2) lemma Int_UNIV [simp]: "A \ B = UNIV \ A = UNIV \ B = UNIV" by (fact inf_eq_top_iff) (* already simp *) lemma Int_subset_iff [simp]: "C \ A \ B \ C \ A \ C \ B" by (fact le_inf_iff) lemma Int_Collect: "x \ A \ {x. P x} \ x \ A \ P x" by blast text \\<^medskip> \Un\.\ lemma Un_absorb: "A \ A = A" by (fact sup_idem) (* already simp *) lemma Un_left_absorb: "A \ (A \ B) = A \ B" by (fact sup_left_idem) lemma Un_commute: "A \ B = B \ A" by (fact sup_commute) lemma Un_left_commute: "A \ (B \ C) = B \ (A \ C)" by (fact sup_left_commute) lemma Un_assoc: "(A \ B) \ C = A \ (B \ C)" by (fact sup_assoc) lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute \ \Union is an AC-operator\ lemma Un_absorb1: "A \ B \ A \ B = B" by (fact sup_absorb2) lemma Un_absorb2: "B \ A \ A \ B = A" by (fact sup_absorb1) lemma Un_empty_left: "{} \ B = B" by (fact sup_bot_left) (* already simp *) lemma Un_empty_right: "A \ {} = A" by (fact sup_bot_right) (* already simp *) lemma Un_UNIV_left: "UNIV \ B = UNIV" by (fact sup_top_left) (* already simp *) lemma Un_UNIV_right: "A \ UNIV = UNIV" by (fact sup_top_right) (* already simp *) lemma Un_insert_left [simp]: "(insert a B) \ C = insert a (B \ C)" by blast lemma Un_insert_right [simp]: "A \ (insert a B) = insert a (A \ B)" by blast lemma Int_insert_left: "(insert a B) \ C = (if a \ C then insert a (B \ C) else B \ C)" by auto lemma Int_insert_left_if0 [simp]: "a \ C \ (insert a B) \ C = B \ C" by auto lemma Int_insert_left_if1 [simp]: "a \ C \ (insert a B) \ C = insert a (B \ C)" by auto lemma Int_insert_right: "A \ (insert a B) = (if a \ A then insert a (A \ B) else A \ B)" by auto lemma Int_insert_right_if0 [simp]: "a \ A \ A \ (insert a B) = A \ B" by auto lemma Int_insert_right_if1 [simp]: "a \ A \ A \ (insert a B) = insert a (A \ B)" by auto lemma Un_Int_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by (fact sup_inf_distrib1) lemma Un_Int_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by (fact sup_inf_distrib2) lemma Un_Int_crazy: "(A \ B) \ (B \ C) \ (C \ A) = (A \ B) \ (B \ C) \ (C \ A)" by blast lemma subset_Un_eq: "A \ B \ A \ B = B" by (fact le_iff_sup) lemma Un_empty [iff]: "A \ B = {} \ A = {} \ B = {}" by (fact sup_eq_bot_iff) (* FIXME: already simp *) lemma Un_subset_iff [simp]: "A \ B \ C \ A \ C \ B \ C" by (fact le_sup_iff) lemma Un_Diff_Int: "(A - B) \ (A \ B) = A" by blast lemma Diff_Int2: "A \ C - B \ C = A \ C - B" by blast lemma subset_UnE: assumes "C \ A \ B" obtains A' B' where "A' \ A" "B' \ B" "C = A' \ B'" proof show "C \ A \ A" "C \ B \ B" "C = (C \ A) \ (C \ B)" using assms by blast+ qed lemma Un_Int_eq [simp]: "(S \ T) \ S = S" "(S \ T) \ T = T" "S \ (S \ T) = S" "T \ (S \ T) = T" by auto lemma Int_Un_eq [simp]: "(S \ T) \ S = S" "(S \ T) \ T = T" "S \ (S \ T) = S" "T \ (S \ T) = T" by auto text \\<^medskip> Set complement\ lemma Compl_disjoint [simp]: "A \ - A = {}" by (fact inf_compl_bot) lemma Compl_disjoint2 [simp]: "- A \ A = {}" by (fact compl_inf_bot) lemma Compl_partition: "A \ - A = UNIV" by (fact sup_compl_top) lemma Compl_partition2: "- A \ A = UNIV" by (fact compl_sup_top) lemma double_complement: "- (-A) = A" for A :: "'a set" by (fact double_compl) (* already simp *) lemma Compl_Un: "- (A \ B) = (- A) \ (- B)" by (fact compl_sup) (* already simp *) lemma Compl_Int: "- (A \ B) = (- A) \ (- B)" by (fact compl_inf) (* already simp *) lemma subset_Compl_self_eq: "A \ - A \ A = {}" by blast lemma Un_Int_assoc_eq: "(A \ B) \ C = A \ (B \ C) \ C \ A" \ \Halmos, Naive Set Theory, page 16.\ by blast lemma Compl_UNIV_eq: "- UNIV = {}" by (fact compl_top_eq) (* already simp *) lemma Compl_empty_eq: "- {} = UNIV" by (fact compl_bot_eq) (* already simp *) lemma Compl_subset_Compl_iff [iff]: "- A \ - B \ B \ A" by (fact compl_le_compl_iff) (* FIXME: already simp *) lemma Compl_eq_Compl_iff [iff]: "- A = - B \ A = B" for A B :: "'a set" by (fact compl_eq_compl_iff) (* FIXME: already simp *) lemma Compl_insert: "- insert x A = (- A) - {x}" by blast text \\<^medskip> Bounded quantifiers. The following are not added to the default simpset because (a) they duplicate the body and (b) there are no similar rules for \Int\. \ lemma ball_Un: "(\x \ A \ B. P x) \ (\x\A. P x) \ (\x\B. P x)" by blast lemma bex_Un: "(\x \ A \ B. P x) \ (\x\A. P x) \ (\x\B. P x)" by blast text \\<^medskip> Set difference.\ lemma Diff_eq: "A - B = A \ (- B)" by blast lemma Diff_eq_empty_iff [simp]: "A - B = {} \ A \ B" by blast lemma Diff_cancel [simp]: "A - A = {}" by blast lemma Diff_idemp [simp]: "(A - B) - B = A - B" for A B :: "'a set" by blast lemma Diff_triv: "A \ B = {} \ A - B = A" by (blast elim: equalityE) lemma empty_Diff [simp]: "{} - A = {}" by blast lemma Diff_empty [simp]: "A - {} = A" by blast lemma Diff_UNIV [simp]: "A - UNIV = {}" by blast lemma Diff_insert0 [simp]: "x \ A \ A - insert x B = A - B" by blast lemma Diff_insert: "A - insert a B = A - B - {a}" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a 0\\ by blast lemma Diff_insert2: "A - insert a B = A - {a} - B" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a 0\\ by blast lemma insert_Diff_if: "insert x A - B = (if x \ B then A - B else insert x (A - B))" by auto lemma insert_Diff1 [simp]: "x \ B \ insert x A - B = A - B" by blast lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert a A" by blast lemma insert_Diff: "a \ A \ insert a (A - {a}) = A" by blast lemma Diff_insert_absorb: "x \ A \ (insert x A) - {x} = A" by auto lemma Diff_disjoint [simp]: "A \ (B - A) = {}" by blast lemma Diff_partition: "A \ B \ A \ (B - A) = B" by blast lemma double_diff: "A \ B \ B \ C \ B - (C - A) = A" by blast lemma Un_Diff_cancel [simp]: "A \ (B - A) = A \ B" by blast lemma Un_Diff_cancel2 [simp]: "(B - A) \ A = B \ A" by blast lemma Diff_Un: "A - (B \ C) = (A - B) \ (A - C)" by blast lemma Diff_Int: "A - (B \ C) = (A - B) \ (A - C)" by blast lemma Diff_Diff_Int: "A - (A - B) = A \ B" by blast lemma Un_Diff: "(A \ B) - C = (A - C) \ (B - C)" by blast lemma Int_Diff: "(A \ B) - C = A \ (B - C)" by blast lemma Diff_Int_distrib: "C \ (A - B) = (C \ A) - (C \ B)" by blast lemma Diff_Int_distrib2: "(A - B) \ C = (A \ C) - (B \ C)" by blast lemma Diff_Compl [simp]: "A - (- B) = A \ B" by auto lemma Compl_Diff_eq [simp]: "- (A - B) = - A \ B" by blast lemma subset_Compl_singleton [simp]: "A \ - {b} \ b \ A" by blast text \\<^medskip> Quantification over type \<^typ>\bool\.\ lemma bool_induct: "P True \ P False \ P x" by (cases x) auto lemma all_bool_eq: "(\b. P b) \ P True \ P False" by (auto intro: bool_induct) lemma bool_contrapos: "P x \ \ P False \ P True" by (cases x) auto lemma ex_bool_eq: "(\b. P b) \ P True \ P False" by (auto intro: bool_contrapos) lemma UNIV_bool: "UNIV = {False, True}" by (auto intro: bool_induct) text \\<^medskip> \Pow\\ lemma Pow_empty [simp]: "Pow {} = {{}}" by (auto simp add: Pow_def) lemma Pow_singleton_iff [simp]: "Pow X = {Y} \ X = {} \ Y = {}" by blast (* somewhat slow *) lemma Pow_insert: "Pow (insert a A) = Pow A \ (insert a ` Pow A)" by (blast intro: image_eqI [where ?x = "u - {a}" for u]) lemma Pow_Compl: "Pow (- A) = {- B | B. A \ Pow B}" by (blast intro: exI [where ?x = "- u" for u]) lemma Pow_UNIV [simp]: "Pow UNIV = UNIV" by blast lemma Un_Pow_subset: "Pow A \ Pow B \ Pow (A \ B)" by blast lemma Pow_Int_eq [simp]: "Pow (A \ B) = Pow A \ Pow B" by blast text \\<^medskip> Miscellany.\ lemma set_eq_subset: "A = B \ A \ B \ B \ A" by blast lemma subset_iff: "A \ B \ (\t. t \ A \ t \ B)" by blast lemma subset_iff_psubset_eq: "A \ B \ A \ B \ A = B" unfolding less_le by blast lemma all_not_in_conv [simp]: "(\x. x \ A) \ A = {}" by blast lemma ex_in_conv: "(\x. x \ A) \ A \ {}" by blast lemma ball_simps [simp, no_atp]: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\P. (\x\{}. P x) \ True" "\P. (\x\UNIV. P x) \ (\x. P x)" "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. P x))" "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" "\A P. (\ (\x\A. P x)) \ (\x\A. \ P x)" by auto lemma bex_simps [simp, no_atp]: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\P. (\x\{}. P x) \ False" "\P. (\x\UNIV. P x) \ (\x. P x)" "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. P x))" "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" "\A P. (\(\x\A. P x)) \ (\x\A. \ P x)" by auto lemma ex_image_cong_iff [simp, no_atp]: "(\x. x\f`A) \ A \ {}" "(\x. x\f`A \ P x) \ (\x\A. P (f x))" by auto subsubsection \Monotonicity of various operations\ lemma image_mono: "A \ B \ f ` A \ f ` B" by blast lemma Pow_mono: "A \ B \ Pow A \ Pow B" by blast lemma insert_mono: "C \ D \ insert a C \ insert a D" by blast lemma Un_mono: "A \ C \ B \ D \ A \ B \ C \ D" by (fact sup_mono) lemma Int_mono: "A \ C \ B \ D \ A \ B \ C \ D" by (fact inf_mono) lemma Diff_mono: "A \ C \ D \ B \ A - B \ C - D" by blast lemma Compl_anti_mono: "A \ B \ - B \ - A" by (fact compl_mono) text \\<^medskip> Monotonicity of implications.\ lemma in_mono: "A \ B \ x \ A \ x \ B" by (rule impI) (erule subsetD) lemma conj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma disj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma imp_mono: "Q1 \ P1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma imp_refl: "P \ P" .. lemma not_mono: "Q \ P \ \ P \ \ Q" by iprover lemma ex_mono: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover lemma all_mono: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover lemma Collect_mono: "(\x. P x \ Q x) \ Collect P \ Collect Q" by blast lemma Int_Collect_mono: "A \ B \ (\x. x \ A \ P x \ Q x) \ A \ Collect P \ B \ Collect Q" by blast lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono Collect_mono in_mono lemma eq_to_mono: "a = b \ c = d \ b \ d \ a \ c" by iprover subsubsection \Inverse image of a function\ definition vimage :: "('a \ 'b) \ 'b set \ 'a set" (infixr "-`" 90) where "f -` B \ {x. f x \ B}" lemma vimage_eq [simp]: "a \ f -` B \ f a \ B" unfolding vimage_def by blast lemma vimage_singleton_eq: "a \ f -` {b} \ f a = b" by simp lemma vimageI [intro]: "f a = b \ b \ B \ a \ f -` B" unfolding vimage_def by blast lemma vimageI2: "f a \ A \ a \ f -` A" unfolding vimage_def by fast lemma vimageE [elim!]: "a \ f -` B \ (\x. f a = x \ x \ B \ P) \ P" unfolding vimage_def by blast lemma vimageD: "a \ f -` A \ f a \ A" unfolding vimage_def by fast lemma vimage_empty [simp]: "f -` {} = {}" by blast lemma vimage_Compl: "f -` (- A) = - (f -` A)" by blast lemma vimage_Un [simp]: "f -` (A \ B) = (f -` A) \ (f -` B)" by blast lemma vimage_Int [simp]: "f -` (A \ B) = (f -` A) \ (f -` B)" by fast lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}" by blast lemma vimage_Collect: "(\x. P (f x) = Q x) \ f -` (Collect P) = Collect Q" by blast lemma vimage_insert: "f -` (insert a B) = (f -` {a}) \ (f -` B)" \ \NOT suitable for rewriting because of the recurrence of \{a}\.\ by blast lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)" by blast lemma vimage_UNIV [simp]: "f -` UNIV = UNIV" by blast lemma vimage_mono: "A \ B \ f -` A \ f -` B" \ \monotonicity\ by blast lemma vimage_image_eq: "f -` (f ` A) = {y. \x\A. f x = f y}" by (blast intro: sym) lemma image_vimage_subset: "f ` (f -` A) \ A" by blast lemma image_vimage_eq [simp]: "f ` (f -` A) = A \ range f" by blast lemma image_subset_iff_subset_vimage: "f ` A \ B \ A \ f -` B" by blast lemma vimage_const [simp]: "((\x. c) -` A) = (if c \ A then UNIV else {})" by auto lemma vimage_if [simp]: "((\x. if x \ B then c else d) -` A) = (if c \ A then (if d \ A then UNIV else B) else if d \ A then - B else {})" by (auto simp add: vimage_def) lemma vimage_inter_cong: "(\ w. w \ S \ f w = g w) \ f -` y \ S = g -` y \ S" by auto lemma vimage_ident [simp]: "(\x. x) -` Y = Y" by blast subsubsection \Singleton sets\ definition is_singleton :: "'a set \ bool" where "is_singleton A \ (\x. A = {x})" lemma is_singletonI [simp, intro!]: "is_singleton {x}" unfolding is_singleton_def by simp lemma is_singletonI': "A \ {} \ (\x y. x \ A \ y \ A \ x = y) \ is_singleton A" unfolding is_singleton_def by blast lemma is_singletonE: "is_singleton A \ (\x. A = {x} \ P) \ P" unfolding is_singleton_def by blast subsubsection \Getting the contents of a singleton set\ definition the_elem :: "'a set \ 'a" where "the_elem X = (THE x. X = {x})" lemma the_elem_eq [simp]: "the_elem {x} = x" by (simp add: the_elem_def) lemma is_singleton_the_elem: "is_singleton A \ A = {the_elem A}" by (auto simp: is_singleton_def) lemma the_elem_image_unique: assumes "A \ {}" and *: "\y. y \ A \ f y = f x" shows "the_elem (f ` A) = f x" unfolding the_elem_def proof (rule the1_equality) from \A \ {}\ obtain y where "y \ A" by auto with * have "f x = f y" by simp with \y \ A\ have "f x \ f ` A" by blast with * show "f ` A = {f x}" by auto then show "\!x. f ` A = {x}" by auto qed subsubsection \Least value operator\ lemma Least_mono: "mono f \ \x\S. \y\S. x \ y \ (LEAST y. y \ f ` S) = f (LEAST x. x \ S)" for f :: "'a::order \ 'b::order" \ \Courtesy of Stephan Merz\ apply clarify apply (erule_tac P = "\x. x \ S" in LeastI2_order) apply fast apply (rule LeastI2_order) apply (auto elim: monoD intro!: order_antisym) done subsubsection \Monad operation\ definition bind :: "'a set \ ('a \ 'b set) \ 'b set" where "bind A f = {x. \B \ f`A. x \ B}" hide_const (open) bind lemma bind_bind: "Set.bind (Set.bind A B) C = Set.bind A (\x. Set.bind (B x) C)" for A :: "'a set" by (auto simp: bind_def) lemma empty_bind [simp]: "Set.bind {} f = {}" by (simp add: bind_def) lemma nonempty_bind_const: "A \ {} \ Set.bind A (\_. B) = B" by (auto simp: bind_def) lemma bind_const: "Set.bind A (\_. B) = (if A = {} then {} else B)" by (auto simp: bind_def) lemma bind_singleton_conv_image: "Set.bind A (\x. {f x}) = f ` A" by (auto simp: bind_def) subsubsection \Operations for execution\ definition is_empty :: "'a set \ bool" where [code_abbrev]: "is_empty A \ A = {}" hide_const (open) is_empty definition remove :: "'a \ 'a set \ 'a set" where [code_abbrev]: "remove x A = A - {x}" hide_const (open) remove lemma member_remove [simp]: "x \ Set.remove y A \ x \ A \ x \ y" by (simp add: remove_def) definition filter :: "('a \ bool) \ 'a set \ 'a set" where [code_abbrev]: "filter P A = {a \ A. P a}" hide_const (open) filter lemma member_filter [simp]: "x \ Set.filter P A \ x \ A \ P x" by (simp add: filter_def) instantiation set :: (equal) equal begin definition "HOL.equal A B \ A \ B \ B \ A" instance by standard (auto simp add: equal_set_def) end text \Misc\ definition pairwise :: "('a \ 'a \ bool) \ 'a set \ bool" where "pairwise R S \ (\x \ S. \y \ S. x \ y \ R x y)" lemma pairwise_alt: "pairwise R S \ (\x\S. \y\S-{x}. R x y)" by (auto simp add: pairwise_def) lemma pairwise_trivial [simp]: "pairwise (\i j. j \ i) I" by (auto simp: pairwise_def) lemma pairwiseI [intro?]: "pairwise R S" if "\x y. x \ S \ y \ S \ x \ y \ R x y" using that by (simp add: pairwise_def) lemma pairwiseD: "R x y" and "R y x" if "pairwise R S" "x \ S" and "y \ S" and "x \ y" using that by (simp_all add: pairwise_def) lemma pairwise_empty [simp]: "pairwise P {}" by (simp add: pairwise_def) lemma pairwise_singleton [simp]: "pairwise P {A}" by (simp add: pairwise_def) lemma pairwise_insert: "pairwise r (insert x s) \ (\y. y \ s \ y \ x \ r x y \ r y x) \ pairwise r s" by (force simp: pairwise_def) lemma pairwise_subset: "pairwise P S \ T \ S \ pairwise P T" by (force simp: pairwise_def) lemma pairwise_mono: "\pairwise P A; \x y. P x y \ Q x y; B \ A\ \ pairwise Q B" by (fastforce simp: pairwise_def) lemma pairwise_imageI: "pairwise P (f ` A)" if "\x y. x \ A \ y \ A \ x \ y \ f x \ f y \ P (f x) (f y)" using that by (auto intro: pairwiseI) lemma pairwise_image: "pairwise r (f ` s) \ pairwise (\x y. (f x \ f y) \ r (f x) (f y)) s" by (force simp: pairwise_def) definition disjnt :: "'a set \ 'a set \ bool" where "disjnt A B \ A \ B = {}" lemma disjnt_self_iff_empty [simp]: "disjnt S S \ S = {}" by (auto simp: disjnt_def) lemma disjnt_iff: "disjnt A B \ (\x. \ (x \ A \ x \ B))" by (force simp: disjnt_def) lemma disjnt_sym: "disjnt A B \ disjnt B A" using disjnt_iff by blast lemma disjnt_empty1 [simp]: "disjnt {} A" and disjnt_empty2 [simp]: "disjnt A {}" by (auto simp: disjnt_def) lemma disjnt_insert1 [simp]: "disjnt (insert a X) Y \ a \ Y \ disjnt X Y" by (simp add: disjnt_def) lemma disjnt_insert2 [simp]: "disjnt Y (insert a X) \ a \ Y \ disjnt Y X" by (simp add: disjnt_def) lemma disjnt_subset1 : "\disjnt X Y; Z \ X\ \ disjnt Z Y" by (auto simp: disjnt_def) lemma disjnt_subset2 : "\disjnt X Y; Z \ Y\ \ disjnt X Z" by (auto simp: disjnt_def) lemma disjnt_Un1 [simp]: "disjnt (A \ B) C \ disjnt A C \ disjnt B C" by (auto simp: disjnt_def) lemma disjnt_Un2 [simp]: "disjnt C (A \ B) \ disjnt C A \ disjnt C B" by (auto simp: disjnt_def) lemma disjoint_image_subset: "\pairwise disjnt \; \X. X \ \ \ f X \ X\ \ pairwise disjnt (f `\)" unfolding disjnt_def pairwise_def by fast lemma pairwise_disjnt_iff: "pairwise disjnt \ \ (\x. \\<^sub>\\<^sub>1 X. X \ \ \ x \ X)" by (auto simp: Uniq_def disjnt_iff pairwise_def) +lemma disjnt_insert: \<^marker>\contributor \Lars Hupel\\ + \disjnt (insert x M) N\ if \x \ N\ \disjnt M N\ + using that by (simp add: disjnt_def) + lemma Int_emptyI: "(\x. x \ A \ x \ B \ False) \ A \ B = {}" by blast lemma in_image_insert_iff: assumes "\C. C \ B \ x \ C" shows "A \ insert x ` B \ x \ A \ A - {x} \ B" (is "?P \ ?Q") proof assume ?P then show ?Q using assms by auto next assume ?Q then have "x \ A" and "A - {x} \ B" by simp_all from \A - {x} \ B\ have "insert x (A - {x}) \ insert x ` B" by (rule imageI) also from \x \ A\ have "insert x (A - {x}) = A" by auto finally show ?P . qed hide_const (open) member not_member lemmas equalityI = subset_antisym lemmas set_mp = subsetD lemmas set_rev_mp = rev_subsetD ML \ val Ball_def = @{thm Ball_def} val Bex_def = @{thm Bex_def} val CollectD = @{thm CollectD} val CollectE = @{thm CollectE} val CollectI = @{thm CollectI} val Collect_conj_eq = @{thm Collect_conj_eq} val Collect_mem_eq = @{thm Collect_mem_eq} val IntD1 = @{thm IntD1} val IntD2 = @{thm IntD2} val IntE = @{thm IntE} val IntI = @{thm IntI} val Int_Collect = @{thm Int_Collect} val UNIV_I = @{thm UNIV_I} val UNIV_witness = @{thm UNIV_witness} val UnE = @{thm UnE} val UnI1 = @{thm UnI1} val UnI2 = @{thm UnI2} val ballE = @{thm ballE} val ballI = @{thm ballI} val bexCI = @{thm bexCI} val bexE = @{thm bexE} val bexI = @{thm bexI} val bex_triv = @{thm bex_triv} val bspec = @{thm bspec} val contra_subsetD = @{thm contra_subsetD} val equalityCE = @{thm equalityCE} val equalityD1 = @{thm equalityD1} val equalityD2 = @{thm equalityD2} val equalityE = @{thm equalityE} val equalityI = @{thm equalityI} val imageE = @{thm imageE} val imageI = @{thm imageI} val image_Un = @{thm image_Un} val image_insert = @{thm image_insert} val insert_commute = @{thm insert_commute} val insert_iff = @{thm insert_iff} val mem_Collect_eq = @{thm mem_Collect_eq} val rangeE = @{thm rangeE} val rangeI = @{thm rangeI} val range_eqI = @{thm range_eqI} val subsetCE = @{thm subsetCE} val subsetD = @{thm subsetD} val subsetI = @{thm subsetI} val subset_refl = @{thm subset_refl} val subset_trans = @{thm subset_trans} val vimageD = @{thm vimageD} val vimageE = @{thm vimageE} val vimageI = @{thm vimageI} val vimageI2 = @{thm vimageI2} val vimage_Collect = @{thm vimage_Collect} val vimage_Int = @{thm vimage_Int} val vimage_Un = @{thm vimage_Un} \ end