diff --git a/src/HOL/Fun.thy b/src/HOL/Fun.thy --- a/src/HOL/Fun.thy +++ b/src/HOL/Fun.thy @@ -1,1108 +1,1120 @@ (* 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 +subsection \Transpositions\ + +lemma swap_id_idempotent [simp]: "Fun.swap a b id \ Fun.swap a b id = id" + by (rule ext) (auto simp add: Fun.swap_def) + +lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)" + by (simp add: Fun.swap_def) + +lemma bij_swap_compose_bij: + \bij (Fun.swap a b id \ p)\ if \bij p\ + using that by (rule bij_comp) 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/Hilbert_Choice.thy b/src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy +++ b/src/HOL/Hilbert_Choice.thy @@ -1,1222 +1,1231 @@ (* Title: HOL/Hilbert_Choice.thy Author: Lawrence C Paulson, Tobias Nipkow Author: Viorel Preoteasa (Results about complete distributive lattices) Copyright 2001 University of Cambridge *) section \Hilbert's Epsilon-Operator and the Axiom of Choice\ theory Hilbert_Choice imports Wellfounded keywords "specification" :: thy_goal_defn begin subsection \Hilbert's epsilon\ axiomatization Eps :: "('a \ bool) \ 'a" where someI: "P x \ P (Eps P)" syntax (epsilon) "_Eps" :: "pttrn \ bool \ 'a" ("(3\_./ _)" [0, 10] 10) syntax (input) "_Eps" :: "pttrn \ bool \ 'a" ("(3@ _./ _)" [0, 10] 10) syntax "_Eps" :: "pttrn \ bool \ 'a" ("(3SOME _./ _)" [0, 10] 10) translations "SOME x. P" \ "CONST Eps (\x. P)" print_translation \ [(\<^const_syntax>\Eps\, fn _ => fn [Abs abs] => let val (x, t) = Syntax_Trans.atomic_abs_tr' abs in Syntax.const \<^syntax_const>\_Eps\ $ x $ t end)] \ \ \to avoid eta-contraction of body\ definition inv_into :: "'a set \ ('a \ 'b) \ ('b \ 'a)" where "inv_into A f = (\x. SOME y. y \ A \ f y = x)" lemma inv_into_def2: "inv_into A f x = (SOME y. y \ A \ f y = x)" by(simp add: inv_into_def) abbreviation inv :: "('a \ 'b) \ ('b \ 'a)" where "inv \ inv_into UNIV" subsection \Hilbert's Epsilon-operator\ lemma Eps_cong: assumes "\x. P x = Q x" shows "Eps P = Eps Q" using ext[of P Q, OF assms] by simp text \ Easier to use than \someI\ if the witness comes from an existential formula. \ lemma someI_ex [elim?]: "\x. P x \ P (SOME x. P x)" by (elim exE someI) lemma some_eq_imp: assumes "Eps P = a" "P b" shows "P a" using assms someI_ex by force text \ Easier to use than \someI\ because the conclusion has only one occurrence of \<^term>\P\. \ lemma someI2: "P a \ (\x. P x \ Q x) \ Q (SOME x. P x)" by (blast intro: someI) text \ Easier to use than \someI2\ if the witness comes from an existential formula. \ lemma someI2_ex: "\a. P a \ (\x. P x \ Q x) \ Q (SOME x. P x)" by (blast intro: someI2) lemma someI2_bex: "\a\A. P a \ (\x. x \ A \ P x \ Q x) \ Q (SOME x. x \ A \ P x)" by (blast intro: someI2) lemma some_equality [intro]: "P a \ (\x. P x \ x = a) \ (SOME x. P x) = a" by (blast intro: someI2) lemma some1_equality: "\!x. P x \ P a \ (SOME x. P x) = a" by blast lemma some_eq_ex: "P (SOME x. P x) \ (\x. P x)" by (blast intro: someI) lemma some_in_eq: "(SOME x. x \ A) \ A \ A \ {}" unfolding ex_in_conv[symmetric] by (rule some_eq_ex) lemma some_eq_trivial [simp]: "(SOME y. y = x) = x" by (rule some_equality) (rule refl) lemma some_sym_eq_trivial [simp]: "(SOME y. x = y) = x" by (iprover intro: some_equality) subsection \Axiom of Choice, Proved Using the Description Operator\ lemma choice: "\x. \y. Q x y \ \f. \x. Q x (f x)" by (fast elim: someI) lemma bchoice: "\x\S. \y. Q x y \ \f. \x\S. Q x (f x)" by (fast elim: someI) lemma choice_iff: "(\x. \y. Q x y) \ (\f. \x. Q x (f x))" by (fast elim: someI) lemma choice_iff': "(\x. P x \ (\y. Q x y)) \ (\f. \x. P x \ Q x (f x))" by (fast elim: someI) lemma bchoice_iff: "(\x\S. \y. Q x y) \ (\f. \x\S. Q x (f x))" by (fast elim: someI) lemma bchoice_iff': "(\x\S. P x \ (\y. Q x y)) \ (\f. \x\S. P x \ Q x (f x))" by (fast elim: someI) lemma dependent_nat_choice: assumes 1: "\x. P 0 x" and 2: "\x n. P n x \ \y. P (Suc n) y \ Q n x y" shows "\f. \n. P n (f n) \ Q n (f n) (f (Suc n))" proof (intro exI allI conjI) fix n define f where "f = rec_nat (SOME x. P 0 x) (\n x. SOME y. P (Suc n) y \ Q n x y)" then have "P 0 (f 0)" "\n. P n (f n) \ P (Suc n) (f (Suc n)) \ Q n (f n) (f (Suc n))" using someI_ex[OF 1] someI_ex[OF 2] by simp_all then show "P n (f n)" "Q n (f n) (f (Suc n))" by (induct n) auto qed lemma finite_subset_Union: assumes "finite A" "A \ \\" obtains \ where "finite \" "\ \ \" "A \ \\" proof - have "\x\A. \B\\. x\B" using assms by blast then obtain f where f: "\x. x \ A \ f x \ \ \ x \ f x" by (auto simp add: bchoice_iff Bex_def) show thesis proof show "finite (f ` A)" using assms by auto qed (use f in auto) qed subsection \Function Inverse\ lemma inv_def: "inv f = (\y. SOME x. f x = y)" by (simp add: inv_into_def) lemma inv_into_into: "x \ f ` A \ inv_into A f x \ A" by (simp add: inv_into_def) (fast intro: someI2) lemma inv_identity [simp]: "inv (\a. a) = (\a. a)" by (simp add: inv_def) lemma inv_id [simp]: "inv id = id" by (simp add: id_def) lemma inv_into_f_f [simp]: "inj_on f A \ x \ A \ inv_into A f (f x) = x" by (simp add: inv_into_def inj_on_def) (blast intro: someI2) lemma inv_f_f: "inj f \ inv f (f x) = x" by simp lemma f_inv_into_f: "y \ f`A \ f (inv_into A f y) = y" by (simp add: inv_into_def) (fast intro: someI2) lemma inv_into_f_eq: "inj_on f A \ x \ A \ f x = y \ inv_into A f y = x" by (erule subst) (fast intro: inv_into_f_f) lemma inv_f_eq: "inj f \ f x = y \ inv f y = x" by (simp add:inv_into_f_eq) lemma inj_imp_inv_eq: "inj f \ \x. f (g x) = x \ inv f = g" by (blast intro: inv_into_f_eq) text \But is it useful?\ lemma inj_transfer: assumes inj: "inj f" and minor: "\y. y \ range f \ P (inv f y)" shows "P x" proof - have "f x \ range f" by auto then have "P(inv f (f x))" by (rule minor) then show "P x" by (simp add: inv_into_f_f [OF inj]) qed lemma inj_iff: "inj f \ inv f \ f = id" by (simp add: o_def fun_eq_iff) (blast intro: inj_on_inverseI inv_into_f_f) lemma inv_o_cancel[simp]: "inj f \ inv f \ f = id" by (simp add: inj_iff) lemma o_inv_o_cancel[simp]: "inj f \ g \ inv f \ f = g" by (simp add: comp_assoc) lemma inv_into_image_cancel[simp]: "inj_on f A \ S \ A \ inv_into A f ` f ` S = S" by (fastforce simp: image_def) lemma inj_imp_surj_inv: "inj f \ surj (inv f)" by (blast intro!: surjI inv_into_f_f) lemma surj_f_inv_f: "surj f \ f (inv f y) = y" by (simp add: f_inv_into_f) lemma bij_inv_eq_iff: "bij p \ x = inv p y \ p x = y" using surj_f_inv_f[of p] by (auto simp add: bij_def) lemma inv_into_injective: assumes eq: "inv_into A f x = inv_into A f y" and x: "x \ f`A" and y: "y \ f`A" shows "x = y" proof - from eq have "f (inv_into A f x) = f (inv_into A f y)" by simp with x y show ?thesis by (simp add: f_inv_into_f) qed lemma inj_on_inv_into: "B \ f`A \ inj_on (inv_into A f) B" by (blast intro: inj_onI dest: inv_into_injective injD) lemma inj_imp_bij_betw_inv: "inj f \ bij_betw (inv f) (f ` M) M" by (simp add: bij_betw_def image_subsetI inj_on_inv_into) lemma bij_betw_inv_into: "bij_betw f A B \ bij_betw (inv_into A f) B A" by (auto simp add: bij_betw_def inj_on_inv_into) lemma surj_imp_inj_inv: "surj f \ inj (inv f)" by (simp add: inj_on_inv_into) lemma surj_iff: "surj f \ f \ inv f = id" by (auto intro!: surjI simp: surj_f_inv_f fun_eq_iff[where 'b='a]) lemma surj_iff_all: "surj f \ (\x. f (inv f x) = x)" by (simp add: o_def surj_iff fun_eq_iff) lemma surj_imp_inv_eq: assumes "surj f" and gf: "\x. g (f x) = x" shows "inv f = g" proof (rule ext) fix x have "g (f (inv f x)) = inv f x" by (rule gf) then show "inv f x = g x" by (simp add: surj_f_inv_f \surj f\) qed lemma bij_imp_bij_inv: "bij f \ bij (inv f)" by (simp add: bij_def inj_imp_surj_inv surj_imp_inj_inv) lemma inv_equality: "(\x. g (f x) = x) \ (\y. f (g y) = y) \ inv f = g" by (rule ext) (auto simp add: inv_into_def) lemma inv_inv_eq: "bij f \ inv (inv f) = f" by (rule inv_equality) (auto simp add: bij_def surj_f_inv_f) text \ \bij (inv f)\ implies little about \f\. Consider \f :: bool \ bool\ such that \f True = f False = True\. Then it ia consistent with axiom \someI\ that \inv f\ could be any function at all, including the identity function. If \inv f = id\ then \inv f\ is a bijection, but \inj f\, \surj f\ and \inv (inv f) = f\ all fail. \ lemma inv_into_comp: "inj_on f (g ` A) \ inj_on g A \ x \ f ` g ` A \ inv_into A (f \ g) x = (inv_into A g \ inv_into (g ` A) f) x" by (auto simp: f_inv_into_f inv_into_into intro: inv_into_f_eq comp_inj_on) lemma o_inv_distrib: "bij f \ bij g \ inv (f \ g) = inv g \ inv f" by (rule inv_equality) (auto simp add: bij_def surj_f_inv_f) lemma image_f_inv_f: "surj f \ f ` (inv f ` A) = A" by (simp add: surj_f_inv_f image_comp comp_def) lemma image_inv_f_f: "inj f \ inv f ` (f ` A) = A" by simp lemma bij_image_Collect_eq: assumes "bij f" shows "f ` Collect P = {y. P (inv f y)}" proof show "f ` Collect P \ {y. P (inv f y)}" using assms by (force simp add: bij_is_inj) show "{y. P (inv f y)} \ f ` Collect P" using assms by (blast intro: bij_is_surj [THEN surj_f_inv_f, symmetric]) qed lemma bij_vimage_eq_inv_image: assumes "bij f" shows "f -` A = inv f ` A" proof show "f -` A \ inv f ` A" using assms by (blast intro: bij_is_inj [THEN inv_into_f_f, symmetric]) show "inv f ` A \ f -` A" using assms by (auto simp add: bij_is_surj [THEN surj_f_inv_f]) qed lemma inv_fn_o_fn_is_id: fixes f::"'a \ 'a" assumes "bij f" shows "((inv f)^^n) o (f^^n) = (\x. x)" proof - have "((inv f)^^n)((f^^n) x) = x" for x n proof (induction n) case (Suc n) have *: "(inv f) (f y) = y" for y by (simp add: assms bij_is_inj) have "(inv f ^^ Suc n) ((f ^^ Suc n) x) = (inv f^^n) (inv f (f ((f^^n) x)))" by (simp add: funpow_swap1) also have "... = (inv f^^n) ((f^^n) x)" using * by auto also have "... = x" using Suc.IH by auto finally show ?case by simp qed (auto) then show ?thesis unfolding o_def by blast qed lemma fn_o_inv_fn_is_id: fixes f::"'a \ 'a" assumes "bij f" shows "(f^^n) o ((inv f)^^n) = (\x. x)" proof - have "(f^^n) (((inv f)^^n) x) = x" for x n proof (induction n) case (Suc n) have *: "f(inv f y) = y" for y using bij_inv_eq_iff[OF assms] by auto have "(f ^^ Suc n) ((inv f ^^ Suc n) x) = (f^^n) (f (inv f ((inv f^^n) x)))" by (simp add: funpow_swap1) also have "... = (f^^n) ((inv f^^n) x)" using * by auto also have "... = x" using Suc.IH by auto finally show ?case by simp qed (auto) then show ?thesis unfolding o_def by blast qed lemma inv_fn: fixes f::"'a \ 'a" assumes "bij f" shows "inv (f^^n) = ((inv f)^^n)" proof - have "inv (f^^n) x = ((inv f)^^n) x" for x proof (rule inv_into_f_eq) show "inj (f ^^ n)" by (simp add: inj_fn[OF bij_is_inj [OF assms]]) show "(f ^^ n) ((inv f ^^ n) x) = x" using fn_o_inv_fn_is_id[OF assms, THEN fun_cong] by force qed auto then show ?thesis by auto qed lemma mono_inv: fixes f::"'a::linorder \ 'b::linorder" assumes "mono f" "bij f" shows "mono (inv f)" proof fix x y::'b assume "x \ y" from \bij f\ obtain a b where x: "x = f a" and y: "y = f b" by(fastforce simp: bij_def surj_def) show "inv f x \ inv f y" proof (rule le_cases) assume "a \ b" thus ?thesis using \bij f\ x y by(simp add: bij_def inv_f_f) next assume "b \ a" hence "f b \ f a" by(rule monoD[OF \mono f\]) hence "y \ x" using x y by simp hence "x = y" using \x \ y\ by auto thus ?thesis by simp qed qed lemma strict_mono_inv_on_range: fixes f :: "'a::linorder \ 'b::order" assumes "strict_mono f" shows "strict_mono_on (inv f) (range f)" proof (clarsimp simp: strict_mono_on_def) fix x y assume "f x < f y" then show "inv f (f x) < inv f (f y)" using assms strict_mono_imp_inj_on strict_mono_less by fastforce qed lemma mono_bij_Inf: fixes f :: "'a::complete_linorder \ 'b::complete_linorder" assumes "mono f" "bij f" shows "f (Inf A) = Inf (f`A)" proof - have "surj f" using \bij f\ by (auto simp: bij_betw_def) have *: "(inv f) (Inf (f`A)) \ Inf ((inv f)`(f`A))" using mono_Inf[OF mono_inv[OF assms], of "f`A"] by simp have "Inf (f`A) \ f (Inf ((inv f)`(f`A)))" using monoD[OF \mono f\ *] by(simp add: surj_f_inv_f[OF \surj f\]) also have "... = f(Inf A)" using assms by (simp add: bij_is_inj) finally show ?thesis using mono_Inf[OF assms(1), of A] by auto qed lemma finite_fun_UNIVD1: assumes fin: "finite (UNIV :: ('a \ 'b) set)" and card: "card (UNIV :: 'b set) \ Suc 0" shows "finite (UNIV :: 'a set)" proof - let ?UNIV_b = "UNIV :: 'b set" from fin have "finite ?UNIV_b" by (rule finite_fun_UNIVD2) with card have "card ?UNIV_b \ Suc (Suc 0)" by (cases "card ?UNIV_b") (auto simp: card_eq_0_iff) then have "card ?UNIV_b = Suc (Suc (card ?UNIV_b - Suc (Suc 0)))" by simp then obtain b1 b2 :: 'b where b1b2: "b1 \ b2" by (auto simp: card_Suc_eq) from fin have fin': "finite (range (\f :: 'a \ 'b. inv f b1))" by (rule finite_imageI) have "UNIV = range (\f :: 'a \ 'b. inv f b1)" proof (rule UNIV_eq_I) fix x :: 'a from b1b2 have "x = inv (\y. if y = x then b1 else b2) b1" by (simp add: inv_into_def) then show "x \ range (\f::'a \ 'b. inv f b1)" by blast qed with fin' show ?thesis by simp qed text \ Every infinite set contains a countable subset. More precisely we show that a set \S\ is infinite if and only if there exists an injective function from the naturals into \S\. The ``only if'' direction is harder because it requires the construction of a sequence of pairwise different elements of an infinite set \S\. The idea is to construct a sequence of non-empty and infinite subsets of \S\ obtained by successively removing elements of \S\. \ lemma infinite_countable_subset: assumes inf: "\ finite S" shows "\f::nat \ 'a. inj f \ range f \ S" \ \Courtesy of Stephan Merz\ proof - define Sseq where "Sseq = rec_nat S (\n T. T - {SOME e. e \ T})" define pick where "pick n = (SOME e. e \ Sseq n)" for n have *: "Sseq n \ S" "\ finite (Sseq n)" for n by (induct n) (auto simp: Sseq_def inf) then have **: "\n. pick n \ Sseq n" unfolding pick_def by (subst (asm) finite.simps) (auto simp add: ex_in_conv intro: someI_ex) with * have "range pick \ S" by auto moreover have "pick n \ pick (n + Suc m)" for m n proof - have "pick n \ Sseq (n + Suc m)" by (induct m) (auto simp add: Sseq_def pick_def) with ** show ?thesis by auto qed then have "inj pick" by (intro linorder_injI) (auto simp add: less_iff_Suc_add) ultimately show ?thesis by blast qed lemma infinite_iff_countable_subset: "\ finite S \ (\f::nat \ 'a. inj f \ range f \ S)" \ \Courtesy of Stephan Merz\ using finite_imageD finite_subset infinite_UNIV_char_0 infinite_countable_subset by auto lemma image_inv_into_cancel: assumes surj: "f`A = A'" and sub: "B' \ A'" shows "f `((inv_into A f)`B') = B'" using assms proof (auto simp: f_inv_into_f) let ?f' = "inv_into A f" fix a' assume *: "a' \ B'" with sub have "a' \ A'" by auto with surj have "a' = f (?f' a')" by (auto simp: f_inv_into_f) with * show "a' \ f ` (?f' ` B')" by blast qed lemma inv_into_inv_into_eq: assumes "bij_betw f A A'" and a: "a \ A" shows "inv_into A' (inv_into A f) a = f a" proof - let ?f' = "inv_into A f" let ?f'' = "inv_into A' ?f'" from assms have *: "bij_betw ?f' A' A" by (auto simp: bij_betw_inv_into) with a obtain a' where a': "a' \ A'" "?f' a' = a" unfolding bij_betw_def by force with a * have "?f'' a = a'" by (auto simp: f_inv_into_f bij_betw_def) moreover from assms a' have "f a = a'" by (auto simp: bij_betw_def) ultimately show "?f'' a = f a" by simp qed lemma inj_on_iff_surj: assumes "A \ {}" shows "(\f. inj_on f A \ f ` A \ A') \ (\g. g ` A' = A)" proof safe fix f assume inj: "inj_on f A" and incl: "f ` A \ A'" let ?phi = "\a' a. a \ A \ f a = a'" let ?csi = "\a. a \ A" let ?g = "\a'. if a' \ f ` A then (SOME a. ?phi a' a) else (SOME a. ?csi a)" have "?g ` A' = A" proof show "?g ` A' \ A" proof clarify fix a' assume *: "a' \ A'" show "?g a' \ A" proof (cases "a' \ f ` A") case True then obtain a where "?phi a' a" by blast then have "?phi a' (SOME a. ?phi a' a)" using someI[of "?phi a'" a] by blast with True show ?thesis by auto next case False with assms have "?csi (SOME a. ?csi a)" using someI_ex[of ?csi] by blast with False show ?thesis by auto qed qed next show "A \ ?g ` A'" proof - have "?g (f a) = a \ f a \ A'" if a: "a \ A" for a proof - let ?b = "SOME aa. ?phi (f a) aa" from a have "?phi (f a) a" by auto then have *: "?phi (f a) ?b" using someI[of "?phi(f a)" a] by blast then have "?g (f a) = ?b" using a by auto moreover from inj * a have "a = ?b" by (auto simp add: inj_on_def) ultimately have "?g(f a) = a" by simp with incl a show ?thesis by auto qed then show ?thesis by force qed qed then show "\g. g ` A' = A" by blast next fix g let ?f = "inv_into A' g" have "inj_on ?f (g ` A')" by (auto simp: inj_on_inv_into) moreover have "?f (g a') \ A'" if a': "a' \ A'" for a' proof - let ?phi = "\ b'. b' \ A' \ g b' = g a'" from a' have "?phi a'" by auto then have "?phi (SOME b'. ?phi b')" using someI[of ?phi] by blast then show ?thesis by (auto simp: inv_into_def) qed ultimately show "\f. inj_on f (g ` A') \ f ` g ` A' \ A'" by auto qed lemma Ex_inj_on_UNION_Sigma: "\f. (inj_on f (\i \ I. A i) \ f ` (\i \ I. A i) \ (SIGMA i : I. A i))" proof let ?phi = "\a i. i \ I \ a \ A i" let ?sm = "\a. SOME i. ?phi a i" let ?f = "\a. (?sm a, a)" have "inj_on ?f (\i \ I. A i)" by (auto simp: inj_on_def) moreover have "?sm a \ I \ a \ A(?sm a)" if "i \ I" and "a \ A i" for i a using that someI[of "?phi a" i] by auto then have "?f ` (\i \ I. A i) \ (SIGMA i : I. A i)" by auto ultimately show "inj_on ?f (\i \ I. A i) \ ?f ` (\i \ I. A i) \ (SIGMA i : I. A i)" by auto qed lemma inv_unique_comp: assumes fg: "f \ g = id" and gf: "g \ f = id" shows "inv f = g" using fg gf inv_equality[of g f] by (auto simp add: fun_eq_iff) +lemma inv_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id" + by (rule inv_unique_comp) simp_all + +lemma bij_swap_comp: + assumes "bij p" + shows "Fun.swap a b id \ p = Fun.swap (inv p a) (inv p b) p" + using surj_f_inv_f[OF bij_is_surj[OF \bij p\]] + by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF \bij p\]) + lemma subset_image_inj: "S \ f ` T \ (\U. U \ T \ inj_on f U \ S = f ` U)" proof safe show "\U\T. inj_on f U \ S = f ` U" if "S \ f ` T" proof - from that [unfolded subset_image_iff subset_iff] obtain g where g: "\x. x \ S \ g x \ T \ x = f (g x)" by (auto simp add: image_iff Bex_def choice_iff') show ?thesis proof (intro exI conjI) show "g ` S \ T" by (simp add: g image_subsetI) show "inj_on f (g ` S)" using g by (auto simp: inj_on_def) show "S = f ` (g ` S)" using g image_subset_iff by auto qed qed qed blast subsection \Other Consequences of Hilbert's Epsilon\ text \Hilbert's Epsilon and the \<^term>\split\ Operator\ text \Looping simprule!\ lemma split_paired_Eps: "(SOME x. P x) = (SOME (a, b). P (a, b))" by simp lemma Eps_case_prod: "Eps (case_prod P) = (SOME xy. P (fst xy) (snd xy))" by (simp add: split_def) lemma Eps_case_prod_eq [simp]: "(SOME (x', y'). x = x' \ y = y') = (x, y)" by blast text \A relation is wellfounded iff it has no infinite descending chain.\ lemma wf_iff_no_infinite_down_chain: "wf r \ (\f. \i. (f (Suc i), f i) \ r)" (is "_ \ \ ?ex") proof assume "wf r" show "\ ?ex" proof assume ?ex then obtain f where f: "(f (Suc i), f i) \ r" for i by blast from \wf r\ have minimal: "x \ Q \ \z\Q. \y. (y, z) \ r \ y \ Q" for x Q by (auto simp: wf_eq_minimal) let ?Q = "{w. \i. w = f i}" fix n have "f n \ ?Q" by blast from minimal [OF this] obtain j where "(y, f j) \ r \ y \ ?Q" for y by blast with this [OF \(f (Suc j), f j) \ r\] have "f (Suc j) \ ?Q" by simp then show False by blast qed next assume "\ ?ex" then show "wf r" proof (rule contrapos_np) assume "\ wf r" then obtain Q x where x: "x \ Q" and rec: "z \ Q \ \y. (y, z) \ r \ y \ Q" for z by (auto simp add: wf_eq_minimal) obtain descend :: "nat \ 'a" where descend_0: "descend 0 = x" and descend_Suc: "descend (Suc n) = (SOME y. y \ Q \ (y, descend n) \ r)" for n by (rule that [of "rec_nat x (\_ rec. (SOME y. y \ Q \ (y, rec) \ r))"]) simp_all have descend_Q: "descend n \ Q" for n proof (induct n) case 0 with x show ?case by (simp only: descend_0) next case Suc then show ?case by (simp only: descend_Suc) (rule someI2_ex; use rec in blast) qed have "(descend (Suc i), descend i) \ r" for i by (simp only: descend_Suc) (rule someI2_ex; use descend_Q rec in blast) then show "\f. \i. (f (Suc i), f i) \ r" by blast qed qed lemma wf_no_infinite_down_chainE: assumes "wf r" obtains k where "(f (Suc k), f k) \ r" using assms wf_iff_no_infinite_down_chain[of r] by blast text \A dynamically-scoped fact for TFL\ lemma tfl_some: "\P x. P x \ P (Eps P)" by (blast intro: someI) subsection \An aside: bounded accessible part\ text \Finite monotone eventually stable sequences\ lemma finite_mono_remains_stable_implies_strict_prefix: fixes f :: "nat \ 'a::order" assumes S: "finite (range f)" "mono f" and eq: "\n. f n = f (Suc n) \ f (Suc n) = f (Suc (Suc n))" shows "\N. (\n\N. \m\N. m < n \ f m < f n) \ (\n\N. f N = f n)" using assms proof - have "\n. f n = f (Suc n)" proof (rule ccontr) assume "\ ?thesis" then have "\n. f n \ f (Suc n)" by auto with \mono f\ have "\n. f n < f (Suc n)" by (auto simp: le_less mono_iff_le_Suc) with lift_Suc_mono_less_iff[of f] have *: "\n m. n < m \ f n < f m" by auto have "inj f" proof (intro injI) fix x y assume "f x = f y" then show "x = y" by (cases x y rule: linorder_cases) (auto dest: *) qed with \finite (range f)\ have "finite (UNIV::nat set)" by (rule finite_imageD) then show False by simp qed then obtain n where n: "f n = f (Suc n)" .. define N where "N = (LEAST n. f n = f (Suc n))" have N: "f N = f (Suc N)" unfolding N_def using n by (rule LeastI) show ?thesis proof (intro exI[of _ N] conjI allI impI) fix n assume "N \ n" then have "\m. N \ m \ m \ n \ f m = f N" proof (induct rule: dec_induct) case base then show ?case by simp next case (step n) then show ?case using eq [rule_format, of "n - 1"] N by (cases n) (auto simp add: le_Suc_eq) qed from this[of n] \N \ n\ show "f N = f n" by auto next fix n m :: nat assume "m < n" "n \ N" then show "f m < f n" proof (induct rule: less_Suc_induct) case (1 i) then have "i < N" by simp then have "f i \ f (Suc i)" unfolding N_def by (rule not_less_Least) with \mono f\ show ?case by (simp add: mono_iff_le_Suc less_le) next case 2 then show ?case by simp qed qed qed lemma finite_mono_strict_prefix_implies_finite_fixpoint: fixes f :: "nat \ 'a set" assumes S: "\i. f i \ S" "finite S" and ex: "\N. (\n\N. \m\N. m < n \ f m \ f n) \ (\n\N. f N = f n)" shows "f (card S) = (\n. f n)" proof - from ex obtain N where inj: "\n m. n \ N \ m \ N \ m < n \ f m \ f n" and eq: "\n\N. f N = f n" by atomize auto have "i \ N \ i \ card (f i)" for i proof (induct i) case 0 then show ?case by simp next case (Suc i) with inj [of "Suc i" i] have "(f i) \ (f (Suc i))" by auto moreover have "finite (f (Suc i))" using S by (rule finite_subset) ultimately have "card (f i) < card (f (Suc i))" by (intro psubset_card_mono) with Suc inj show ?case by auto qed then have "N \ card (f N)" by simp also have "\ \ card S" using S by (intro card_mono) finally have \
: "f (card S) = f N" using eq by auto moreover have "\ (range f) \ f N" proof clarify fix x n assume "x \ f n" with eq inj [of N] show "x \ f N" by (cases "n < N") (auto simp: not_less) qed ultimately show ?thesis by auto qed subsection \More on injections, bijections, and inverses\ locale bijection = fixes f :: "'a \ 'a" assumes bij: "bij f" begin lemma bij_inv: "bij (inv f)" using bij by (rule bij_imp_bij_inv) lemma surj [simp]: "surj f" using bij by (rule bij_is_surj) lemma inj: "inj f" using bij by (rule bij_is_inj) lemma surj_inv [simp]: "surj (inv f)" using inj by (rule inj_imp_surj_inv) lemma inj_inv: "inj (inv f)" using surj by (rule surj_imp_inj_inv) lemma eqI: "f a = f b \ a = b" using inj by (rule injD) lemma eq_iff [simp]: "f a = f b \ a = b" by (auto intro: eqI) lemma eq_invI: "inv f a = inv f b \ a = b" using inj_inv by (rule injD) lemma eq_inv_iff [simp]: "inv f a = inv f b \ a = b" by (auto intro: eq_invI) lemma inv_left [simp]: "inv f (f a) = a" using inj by (simp add: inv_f_eq) lemma inv_comp_left [simp]: "inv f \ f = id" by (simp add: fun_eq_iff) lemma inv_right [simp]: "f (inv f a) = a" using surj by (simp add: surj_f_inv_f) lemma inv_comp_right [simp]: "f \ inv f = id" by (simp add: fun_eq_iff) lemma inv_left_eq_iff [simp]: "inv f a = b \ f b = a" by auto lemma inv_right_eq_iff [simp]: "b = inv f a \ f b = a" by auto end lemma infinite_imp_bij_betw: assumes infinite: "\ finite A" shows "\h. bij_betw h A (A - {a})" proof (cases "a \ A") case False then have "A - {a} = A" by blast then show ?thesis using bij_betw_id[of A] by auto next case True with infinite have "\ finite (A - {a})" by auto with infinite_iff_countable_subset[of "A - {a}"] obtain f :: "nat \ 'a" where "inj f" and f: "f ` UNIV \ A - {a}" by blast define g where "g n = (if n = 0 then a else f (Suc n))" for n define A' where "A' = g ` UNIV" have *: "\y. f y \ a" using f by blast have 3: "inj_on g UNIV \ g ` UNIV \ A \ a \ g ` UNIV" using \inj f\ f * unfolding inj_on_def g_def by (auto simp add: True image_subset_iff) then have 4: "bij_betw g UNIV A' \ a \ A' \ A' \ A" using inj_on_imp_bij_betw[of g] by (auto simp: A'_def) then have 5: "bij_betw (inv g) A' UNIV" by (auto simp add: bij_betw_inv_into) from 3 obtain n where n: "g n = a" by auto have 6: "bij_betw g (UNIV - {n}) (A' - {a})" by (rule bij_betw_subset) (use 3 4 n in \auto simp: image_set_diff A'_def\) define v where "v m = (if m < n then m else Suc m)" for m have "m < n \ m = n" if "\k. k < n \ m \ Suc k" for m using that [of "m-1"] by auto then have 7: "bij_betw v UNIV (UNIV - {n})" unfolding bij_betw_def inj_on_def v_def by auto define h' where "h' = g \ v \ (inv g)" with 5 6 7 have 8: "bij_betw h' A' (A' - {a})" by (auto simp add: bij_betw_trans) define h where "h b = (if b \ A' then h' b else b)" for b with 8 have "bij_betw h A' (A' - {a})" using bij_betw_cong[of A' h] by auto moreover have "\b \ A - A'. h b = b" by (auto simp: h_def) then have "bij_betw h (A - A') (A - A')" using bij_betw_cong[of "A - A'" h id] bij_betw_id[of "A - A'"] by auto moreover from 4 have "(A' \ (A - A') = {} \ A' \ (A - A') = A) \ ((A' - {a}) \ (A - A') = {} \ (A' - {a}) \ (A - A') = A - {a})" by blast ultimately have "bij_betw h A (A - {a})" using bij_betw_combine[of h A' "A' - {a}" "A - A'" "A - A'"] by simp then show ?thesis by blast qed lemma infinite_imp_bij_betw2: assumes "\ finite A" shows "\h. bij_betw h A (A \ {a})" proof (cases "a \ A") case True then have "A \ {a} = A" by blast then show ?thesis using bij_betw_id[of A] by auto next case False let ?A' = "A \ {a}" from False have "A = ?A' - {a}" by blast moreover from assms have "\ finite ?A'" by auto ultimately obtain f where "bij_betw f ?A' A" using infinite_imp_bij_betw[of ?A' a] by auto then have "bij_betw (inv_into ?A' f) A ?A'" by (rule bij_betw_inv_into) then show ?thesis by auto qed lemma bij_betw_inv_into_left: "bij_betw f A A' \ a \ A \ inv_into A f (f a) = a" unfolding bij_betw_def by clarify (rule inv_into_f_f) lemma bij_betw_inv_into_right: "bij_betw f A A' \ a' \ A' \ f (inv_into A f a') = a'" unfolding bij_betw_def using f_inv_into_f by force lemma bij_betw_inv_into_subset: "bij_betw f A A' \ B \ A \ f ` B = B' \ bij_betw (inv_into A f) B' B" by (auto simp: bij_betw_def intro: inj_on_inv_into) subsection \Specification package -- Hilbertized version\ lemma exE_some: "Ex P \ c \ Eps P \ P c" by (simp only: someI_ex) ML_file \Tools/choice_specification.ML\ subsection \Complete Distributive Lattices -- Properties depending on Hilbert Choice\ context complete_distrib_lattice begin lemma Sup_Inf: "\ (Inf ` A) = \ (Sup ` {f ` A |f. \B\A. f B \ B})" proof (rule antisym) show "\ (Inf ` A) \ \ (Sup ` {f ` A |f. \B\A. f B \ B})" using Inf_lower2 Sup_upper by (fastforce simp add: intro: Sup_least INF_greatest) next show "\ (Sup ` {f ` A |f. \B\A. f B \ B}) \ \ (Inf ` A)" proof (simp add: Inf_Sup, rule SUP_least, simp, safe) fix f assume "\Y. (\f. Y = f ` A \ (\Y\A. f Y \ Y)) \ f Y \ Y" then have B: "\ F . (\ Y \ A . F Y \ Y) \ \ Z \ A . f (F ` A) = F Z" by auto show "\(f ` {f ` A |f. \Y\A. f Y \ Y}) \ \(Inf ` A)" proof (cases "\ Z \ A . \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ Inf Z") case True from this obtain Z where [simp]: "Z \ A" and A: "\(f ` {f ` A |f. \Y\A. f Y \ Y}) \ Inf Z" by blast have B: "... \ \(Inf ` A)" by (simp add: SUP_upper) from A and B show ?thesis by simp next case False then have X: "\ Z . Z \ A \ \ x . x \ Z \ \ \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ x" using Inf_greatest by blast define F where "F = (\ Z . SOME x . x \ Z \ \ \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ x)" have C: "\Y. Y \ A \ F Y \ Y" using X by (simp add: F_def, rule someI2_ex, auto) have E: "\Y. Y \ A \ \ \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ F Y" using X by (simp add: F_def, rule someI2_ex, auto) from C and B obtain Z where D: "Z \ A " and Y: "f (F ` A) = F Z" by blast from E and D have W: "\ \(f ` {f ` A |f. \Y\A. f Y \ Y}) \ F Z" by simp have "\(f ` {f ` A |f. \Y\A. f Y \ Y}) \ f (F ` A)" using C by (blast intro: INF_lower) with W Y show ?thesis by simp qed qed qed lemma dual_complete_distrib_lattice: "class.complete_distrib_lattice Sup Inf sup (\) (>) inf \ \" by (simp add: class.complete_distrib_lattice.intro [OF dual_complete_lattice] class.complete_distrib_lattice_axioms_def Sup_Inf) lemma sup_Inf: "a \ \B = \((\) a ` B)" proof (rule antisym) show "a \ \B \ \((\) a ` B)" using Inf_lower sup.mono by (fastforce intro: INF_greatest) next have "\((\) a ` B) \ \(Sup ` {{f {a}, f B} |f. f {a} = a \ f B \ B})" by (rule INF_greatest, auto simp add: INF_lower) also have "... = \(Inf ` {{a}, B})" by (unfold Sup_Inf, simp) finally show "\((\) a ` B) \ a \ \B" by simp qed lemma inf_Sup: "a \ \B = \((\) a ` B)" using dual_complete_distrib_lattice by (rule complete_distrib_lattice.sup_Inf) lemma INF_SUP: "(\y. \x. P x y) = (\f. \x. P (f x) x)" proof (rule antisym) show "(SUP x. INF y. P (x y) y) \ (INF y. SUP x. P x y)" by (rule SUP_least, rule INF_greatest, rule SUP_upper2, simp_all, rule INF_lower2, simp, blast) next have "(INF y. SUP x. ((P x y))) \ Inf (Sup ` {{P x y | x . True} | y . True })" (is "?A \ ?B") proof (rule INF_greatest, clarsimp) fix y have "?A \ (SUP x. P x y)" by (rule INF_lower, simp) also have "... \ Sup {uu. \x. uu = P x y}" by (simp add: full_SetCompr_eq) finally show "?A \ Sup {uu. \x. uu = P x y}" by simp qed also have "... \ (SUP x. INF y. P (x y) y)" proof (subst Inf_Sup, rule SUP_least, clarsimp) fix f assume A: "\Y. (\y. Y = {uu. \x. uu = P x y}) \ f Y \ Y" have " \(f ` {uu. \y. uu = {uu. \x. uu = P x y}}) \ (\y. P (SOME x. f {P x y |x. True} = P x y) y)" proof (rule INF_greatest, clarsimp) fix y have "(INF x\{uu. \y. uu = {uu. \x. uu = P x y}}. f x) \ f {uu. \x. uu = P x y}" by (rule INF_lower, blast) also have "... \ P (SOME x. f {uu . \x. uu = P x y} = P x y) y" by (rule someI2_ex) (use A in auto) finally show "\(f ` {uu. \y. uu = {uu. \x. uu = P x y}}) \ P (SOME x. f {uu. \x. uu = P x y} = P x y) y" by simp qed also have "... \ (SUP x. INF y. P (x y) y)" by (rule SUP_upper, simp) finally show "\(f ` {uu. \y. uu = {uu. \x. uu = P x y}}) \ (\x. \y. P (x y) y)" by simp qed finally show "(INF y. SUP x. P x y) \ (SUP x. INF y. P (x y) y)" by simp qed lemma INF_SUP_set: "(\B\A. \(g ` B)) = (\B\{f ` A |f. \C\A. f C \ C}. \(g ` B))" (is "_ = (\B\?F. _)") proof (rule antisym) have "\ ((g \ f) ` A) \ \ (g ` B)" if "\B. B \ A \ f B \ B" "B \ A" for f B using that by (auto intro: SUP_upper2 INF_lower2) then show "(\x\?F. \a\x. g a) \ (\x\A. \a\x. g a)" by (auto intro!: SUP_least INF_greatest simp add: image_comp) next show "(\x\A. \a\x. g a) \ (\x\?F. \a\x. g a)" proof (cases "{} \ A") case True then show ?thesis by (rule INF_lower2) simp_all next case False {fix x have "(\x\A. \x\x. g x) \ (\u. if x \ A then if u \ x then g u else \ else \)" proof (cases "x \ A") case True then show ?thesis by (intro INF_lower2 SUP_least SUP_upper2) auto qed auto } then have "(\Y\A. \a\Y. g a) \ (\Y. \y. if Y \ A then if y \ Y then g y else \ else \)" by (rule INF_greatest) also have "... = (\x. \Y. if Y \ A then if x Y \ Y then g (x Y) else \ else \)" by (simp only: INF_SUP) also have "... \ (\x\?F. \a\x. g a)" proof (rule SUP_least) show "(\B. if B \ A then if x B \ B then g (x B) else \ else \) \ (\x\?F. \x\x. g x)" for x proof - define G where "G \ \Y. if x Y \ Y then x Y else (SOME x. x \Y)" have "\Y\A. G Y \ Y" using False some_in_eq G_def by auto then have A: "G ` A \ ?F" by blast show "(\Y. if Y \ A then if x Y \ Y then g (x Y) else \ else \) \ (\x\?F. \x\x. g x)" by (fastforce simp: G_def intro: SUP_upper2 [OF A] INF_greatest INF_lower2) qed qed finally show ?thesis by simp qed qed lemma SUP_INF: "(\y. \x. P x y) = (\x. \y. P (x y) y)" using dual_complete_distrib_lattice by (rule complete_distrib_lattice.INF_SUP) lemma SUP_INF_set: "(\x\A. \ (g ` x)) = (\x\{f ` A |f. \Y\A. f Y \ Y}. \ (g ` x))" using dual_complete_distrib_lattice by (rule complete_distrib_lattice.INF_SUP_set) end (*properties of the former complete_distrib_lattice*) context complete_distrib_lattice begin lemma sup_INF: "a \ (\b\B. f b) = (\b\B. a \ f b)" by (simp add: sup_Inf image_comp) lemma inf_SUP: "a \ (\b\B. f b) = (\b\B. a \ f b)" by (simp add: inf_Sup image_comp) lemma Inf_sup: "\B \ a = (\b\B. b \ a)" by (simp add: sup_Inf sup_commute) lemma Sup_inf: "\B \ a = (\b\B. b \ a)" by (simp add: inf_Sup inf_commute) lemma INF_sup: "(\b\B. f b) \ a = (\b\B. f b \ a)" by (simp add: sup_INF sup_commute) lemma SUP_inf: "(\b\B. f b) \ a = (\b\B. f b \ a)" by (simp add: inf_SUP inf_commute) lemma Inf_sup_eq_top_iff: "(\B \ a = \) \ (\b\B. b \ a = \)" by (simp only: Inf_sup INF_top_conv) lemma Sup_inf_eq_bot_iff: "(\B \ a = \) \ (\b\B. b \ a = \)" by (simp only: Sup_inf SUP_bot_conv) lemma INF_sup_distrib2: "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" by (subst INF_commute) (simp add: sup_INF INF_sup) lemma SUP_inf_distrib2: "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" by (subst SUP_commute) (simp add: inf_SUP SUP_inf) end context complete_boolean_algebra begin lemma dual_complete_boolean_algebra: "class.complete_boolean_algebra Sup Inf sup (\) (>) inf \ \ (\x y. x \ - y) uminus" by (rule class.complete_boolean_algebra.intro, rule dual_complete_distrib_lattice, rule dual_boolean_algebra) end instantiation set :: (type) complete_distrib_lattice begin instance proof (standard, clarsimp) fix A :: "(('a set) set) set" fix x::'a assume A: "\\\A. \X\\. x \ X" define F where "F \ \Y. SOME X. Y \ A \ X \ Y \ x \ X" have "(\S \ F ` A. x \ S)" using A unfolding F_def by (fastforce intro: someI2_ex) moreover have "\Y\A. F Y \ Y" using A unfolding F_def by (fastforce intro: someI2_ex) then have "\f. F ` A = f ` A \ (\Y\A. f Y \ Y)" by blast ultimately show "\X. (\f. X = f ` A \ (\Y\A. f Y \ Y)) \ (\S\X. x \ S)" by auto qed end instance set :: (type) complete_boolean_algebra .. instantiation "fun" :: (type, complete_distrib_lattice) complete_distrib_lattice begin instance by standard (simp add: le_fun_def INF_SUP_set image_comp) end instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra .. context complete_linorder begin subclass complete_distrib_lattice proof (standard, rule ccontr) fix A :: "'a set set" let ?F = "{f ` A |f. \Y\A. f Y \ Y}" assume "\ \(Sup ` A) \ \(Inf ` ?F)" then have C: "\(Sup ` A) > \(Inf ` ?F)" by (simp add: not_le) show False proof (cases "\ z . \(Sup ` A) > z \ z > \(Inf ` ?F)") case True then obtain z where A: "z < \(Sup ` A)" and X: "z > \(Inf ` ?F)" by blast then have B: "\Y. Y \ A \ \k \Y . z < k" using local.less_Sup_iff by(force dest: less_INF_D) define G where "G \ \Y. SOME k . k \ Y \ z < k" have E: "\Y. Y \ A \ G Y \ Y" using B unfolding G_def by (fastforce intro: someI2_ex) have "z \ Inf (G ` A)" proof (rule INF_greatest) show "\Y. Y \ A \ z \ G Y" using B unfolding G_def by (fastforce intro: someI2_ex) qed also have "... \ \(Inf ` ?F)" by (rule SUP_upper) (use E in blast) finally have "z \ \(Inf ` ?F)" by simp with X show ?thesis using local.not_less by blast next case False have B: "\Y. Y \ A \ \ k \Y . \(Inf ` ?F) < k" using C local.less_Sup_iff by(force dest: less_INF_D) define G where "G \ \ Y . SOME k . k \ Y \ \(Inf ` ?F) < k" have E: "\Y. Y \ A \ G Y \ Y" using B unfolding G_def by (fastforce intro: someI2_ex) have "\Y. Y \ A \ \(Sup ` A) \ G Y" using B False local.leI unfolding G_def by (fastforce intro: someI2_ex) then have "\(Sup ` A) \ Inf (G ` A)" by (simp add: local.INF_greatest) also have "Inf (G ` A) \ \(Inf ` ?F)" by (rule SUP_upper) (use E in blast) finally have "\(Sup ` A) \ \(Inf ` ?F)" by simp with C show ?thesis using not_less by blast qed qed end end diff --git a/src/HOL/Library/Permutations.thy b/src/HOL/Library/Permutations.thy --- a/src/HOL/Library/Permutations.thy +++ b/src/HOL/Library/Permutations.thy @@ -1,1499 +1,1472 @@ (* Title: HOL/Library/Permutations.thy Author: Amine Chaieb, University of Cambridge *) section \Permutations, both general and specifically on finite sets.\ theory Permutations imports Multiset Disjoint_Sets begin -subsection \Transpositions\ - -lemma swap_id_idempotent [simp]: "Fun.swap a b id \ Fun.swap a b id = id" - by (rule ext) (auto simp add: Fun.swap_def) - -lemma inv_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id" - by (rule inv_unique_comp) simp_all - -lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)" - by (simp add: Fun.swap_def) - -lemma bij_swap_comp: - assumes "bij p" - shows "Fun.swap a b id \ p = Fun.swap (inv p a) (inv p b) p" - using surj_f_inv_f[OF bij_is_surj[OF \bij p\]] - by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF \bij p\]) - -lemma bij_swap_compose_bij: - assumes "bij p" - shows "bij (Fun.swap a b id \ p)" - by (simp only: bij_swap_comp[OF \bij p\] bij_swap_iff \bij p\) - - -subsection \Basic consequences of the definition\ +subsection \Basic definition and consequences\ definition permutes (infixr "permutes" 41) where "(p permutes S) \ (\x. x \ S \ p x = x) \ (\y. \!x. p x = y)" lemma permutes_in_image: "p permutes S \ p x \ S \ x \ S" unfolding permutes_def by metis lemma permutes_not_in: "f permutes S \ x \ S \ f x = x" by (auto simp: permutes_def) lemma permutes_image: "p permutes S \ p ` S = S" unfolding permutes_def apply (rule set_eqI) apply (simp add: image_iff) apply metis done lemma permutes_inj: "p permutes S \ inj p" unfolding permutes_def inj_def by blast lemma permutes_inj_on: "f permutes S \ inj_on f A" by (auto simp: permutes_def inj_on_def) lemma permutes_surj: "p permutes s \ surj p" unfolding permutes_def surj_def by metis lemma permutes_bij: "p permutes s \ bij p" unfolding bij_def by (metis permutes_inj permutes_surj) lemma permutes_imp_bij: "p permutes S \ bij_betw p S S" by (metis UNIV_I bij_betw_subset permutes_bij permutes_image subsetI) lemma bij_imp_permutes: "bij_betw p S S \ (\x. x \ S \ p x = x) \ p permutes S" unfolding permutes_def bij_betw_def inj_on_def by auto (metis image_iff)+ lemma permutes_inv_o: assumes permutes: "p permutes S" shows "p \ inv p = id" and "inv p \ p = id" using permutes_inj[OF permutes] permutes_surj[OF permutes] unfolding inj_iff[symmetric] surj_iff[symmetric] by blast+ lemma permutes_inverses: fixes p :: "'a \ 'a" assumes permutes: "p permutes S" shows "p (inv p x) = x" and "inv p (p x) = x" using permutes_inv_o[OF permutes, unfolded fun_eq_iff o_def] by auto lemma permutes_subset: "p permutes S \ S \ T \ p permutes T" unfolding permutes_def by blast lemma permutes_empty[simp]: "p permutes {} \ p = id" by (auto simp add: fun_eq_iff permutes_def) lemma permutes_sing[simp]: "p permutes {a} \ p = id" by (simp add: fun_eq_iff permutes_def) metis (*somewhat slow*) lemma permutes_univ: "p permutes UNIV \ (\y. \!x. p x = y)" by (simp add: permutes_def) lemma permutes_inv_eq: "p permutes S \ inv p y = x \ p x = y" unfolding permutes_def inv_def apply auto apply (erule allE[where x=y]) apply (erule allE[where x=y]) apply (rule someI_ex) apply blast apply (rule some1_equality) apply blast apply blast done lemma permutes_swap_id: "a \ S \ b \ S \ Fun.swap a b id permutes S" unfolding permutes_def Fun.swap_def fun_upd_def by auto metis lemma permutes_superset: "p permutes S \ (\x \ S - T. p x = x) \ p permutes T" by (simp add: Ball_def permutes_def) metis lemma permutes_bij_inv_into: \<^marker>\contributor \Lukas Bulwahn\\ fixes A :: "'a set" and B :: "'b set" assumes "p permutes A" and "bij_betw f A B" shows "(\x. if x \ B then f (p (inv_into A f x)) else x) permutes B" proof (rule bij_imp_permutes) from assms have "bij_betw p A A" "bij_betw f A B" "bij_betw (inv_into A f) B A" by (auto simp add: permutes_imp_bij bij_betw_inv_into) then have "bij_betw (f \ p \ inv_into A f) B B" by (simp add: bij_betw_trans) then show "bij_betw (\x. if x \ B then f (p (inv_into A f x)) else x) B B" by (subst bij_betw_cong[where g="f \ p \ inv_into A f"]) auto next fix x assume "x \ B" then show "(if x \ B then f (p (inv_into A f x)) else x) = x" by auto qed lemma permutes_image_mset: \<^marker>\contributor \Lukas Bulwahn\\ assumes "p permutes A" shows "image_mset p (mset_set A) = mset_set A" using assms by (metis image_mset_mset_set bij_betw_imp_inj_on permutes_imp_bij permutes_image) lemma permutes_implies_image_mset_eq: \<^marker>\contributor \Lukas Bulwahn\\ assumes "p permutes A" "\x. x \ A \ f x = f' (p x)" shows "image_mset f' (mset_set A) = image_mset f (mset_set A)" proof - have "f x = f' (p x)" if "x \# mset_set A" for x using assms(2)[of x] that by (cases "finite A") auto with assms have "image_mset f (mset_set A) = image_mset (f' \ p) (mset_set A)" by (auto intro!: image_mset_cong) also have "\ = image_mset f' (image_mset p (mset_set A))" by (simp add: image_mset.compositionality) also have "\ = image_mset f' (mset_set A)" proof - from assms permutes_image_mset have "image_mset p (mset_set A) = mset_set A" by blast then show ?thesis by simp qed finally show ?thesis .. qed subsection \Group properties\ lemma permutes_id: "id permutes S" by (simp add: permutes_def) lemma permutes_compose: "p permutes S \ q permutes S \ q \ p permutes S" unfolding permutes_def o_def by metis lemma permutes_inv: assumes "p permutes S" shows "inv p permutes S" using assms unfolding permutes_def permutes_inv_eq[OF assms] by metis lemma permutes_inv_inv: assumes "p permutes S" shows "inv (inv p) = p" unfolding fun_eq_iff permutes_inv_eq[OF assms] permutes_inv_eq[OF permutes_inv[OF assms]] by blast lemma permutes_invI: assumes perm: "p permutes S" and inv: "\x. x \ S \ p' (p x) = x" and outside: "\x. x \ S \ p' x = x" shows "inv p = p'" proof show "inv p x = p' x" for x proof (cases "x \ S") case True from assms have "p' x = p' (p (inv p x))" by (simp add: permutes_inverses) also from permutes_inv[OF perm] True have "\ = inv p x" by (subst inv) (simp_all add: permutes_in_image) finally show ?thesis .. next case False with permutes_inv[OF perm] show ?thesis by (simp_all add: outside permutes_not_in) qed qed lemma permutes_vimage: "f permutes A \ f -` A = A" by (simp add: bij_vimage_eq_inv_image permutes_bij permutes_image[OF permutes_inv]) subsection \Mapping permutations with bijections\ lemma bij_betw_permutations: assumes "bij_betw f A B" shows "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x) {\. \ permutes A} {\. \ permutes B}" (is "bij_betw ?f _ _") proof - let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)" show ?thesis proof (rule bij_betw_byWitness [of _ ?g], goal_cases) case 3 show ?case using permutes_bij_inv_into[OF _ assms] by auto next case 4 have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) { fix \ assume "\ permutes B" from permutes_bij_inv_into[OF this bij_inv] and assms have "(\x. if x \ A then inv_into A f (\ (f x)) else x) permutes A" by (simp add: inv_into_inv_into_eq cong: if_cong) } from this show ?case by (auto simp: permutes_inv) next case 1 thus ?case using assms by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left dest: bij_betwE) next case 2 moreover have "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) ultimately show ?case using assms by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right dest: bij_betwE) qed qed lemma bij_betw_derangements: assumes "bij_betw f A B" shows "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x) {\. \ permutes A \ (\x\A. \ x \ x)} {\. \ permutes B \ (\x\B. \ x \ x)}" (is "bij_betw ?f _ _") proof - let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)" show ?thesis proof (rule bij_betw_byWitness [of _ ?g], goal_cases) case 3 have "?f \ x \ x" if "\ permutes A" "\x. x \ A \ \ x \ x" "x \ B" for \ x using that and assms by (metis bij_betwE bij_betw_imp_inj_on bij_betw_imp_surj_on inv_into_f_f inv_into_into permutes_imp_bij) with permutes_bij_inv_into[OF _ assms] show ?case by auto next case 4 have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) have "?g \ permutes A" if "\ permutes B" for \ using permutes_bij_inv_into[OF that bij_inv] and assms by (simp add: inv_into_inv_into_eq cong: if_cong) moreover have "?g \ x \ x" if "\ permutes B" "\x. x \ B \ \ x \ x" "x \ A" for \ x using that and assms by (metis bij_betwE bij_betw_imp_surj_on f_inv_into_f permutes_imp_bij) ultimately show ?case by auto next case 1 thus ?case using assms by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left dest: bij_betwE) next case 2 moreover have "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) ultimately show ?case using assms by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right dest: bij_betwE) qed qed subsection \The number of permutations on a finite set\ lemma permutes_insert_lemma: assumes "p permutes (insert a S)" shows "Fun.swap a (p a) id \ p permutes S" apply (rule permutes_superset[where S = "insert a S"]) apply (rule permutes_compose[OF assms]) apply (rule permutes_swap_id, simp) using permutes_in_image[OF assms, of a] apply simp apply (auto simp add: Ball_def Fun.swap_def) done lemma permutes_insert: "{p. p permutes (insert a S)} = (\(b, p). Fun.swap a b id \ p) ` {(b, p). b \ insert a S \ p \ {p. p permutes S}}" proof - have "p permutes insert a S \ (\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S)" for p proof - have "\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S" if p: "p permutes insert a S" proof - let ?b = "p a" let ?q = "Fun.swap a (p a) id \ p" have *: "p = Fun.swap a ?b id \ ?q" by (simp add: fun_eq_iff o_assoc) have **: "?b \ insert a S" unfolding permutes_in_image[OF p] by simp from permutes_insert_lemma[OF p] * ** show ?thesis by blast qed moreover have "p permutes insert a S" if bq: "p = Fun.swap a b id \ q" "b \ insert a S" "q permutes S" for b q proof - from permutes_subset[OF bq(3), of "insert a S"] have q: "q permutes insert a S" by auto have a: "a \ insert a S" by simp from bq(1) permutes_compose[OF q permutes_swap_id[OF a bq(2)]] show ?thesis by simp qed ultimately show ?thesis by blast qed then show ?thesis by auto qed lemma card_permutations: assumes "card S = n" and "finite S" shows "card {p. p permutes S} = fact n" using assms(2,1) proof (induct arbitrary: n) case empty then show ?case by simp next case (insert x F) { fix n assume card_insert: "card (insert x F) = n" let ?xF = "{p. p permutes insert x F}" let ?pF = "{p. p permutes F}" let ?pF' = "{(b, p). b \ insert x F \ p \ ?pF}" let ?g = "(\(b, p). Fun.swap x b id \ p)" have xfgpF': "?xF = ?g ` ?pF'" by (rule permutes_insert[of x F]) from \x \ F\ \finite F\ card_insert have Fs: "card F = n - 1" by auto from \finite F\ insert.hyps Fs have pFs: "card ?pF = fact (n - 1)" by auto then have "finite ?pF" by (auto intro: card_ge_0_finite) with \finite F\ card.insert_remove have pF'f: "finite ?pF'" apply (simp only: Collect_case_prod Collect_mem_eq) apply (rule finite_cartesian_product) apply simp_all done have ginj: "inj_on ?g ?pF'" proof - { fix b p c q assume bp: "(b, p) \ ?pF'" assume cq: "(c, q) \ ?pF'" assume eq: "?g (b, p) = ?g (c, q)" from bp cq have pF: "p permutes F" and qF: "q permutes F" by auto from pF \x \ F\ eq have "b = ?g (b, p) x" by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff) also from qF \x \ F\ eq have "\ = ?g (c, q) x" by (auto simp: swap_def fun_upd_def fun_eq_iff) also from qF \x \ F\ have "\ = c" by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff) finally have "b = c" . then have "Fun.swap x b id = Fun.swap x c id" by simp with eq have "Fun.swap x b id \ p = Fun.swap x b id \ q" by simp then have "Fun.swap x b id \ (Fun.swap x b id \ p) = Fun.swap x b id \ (Fun.swap x b id \ q)" by simp then have "p = q" by (simp add: o_assoc) with \b = c\ have "(b, p) = (c, q)" by simp } then show ?thesis unfolding inj_on_def by blast qed from \x \ F\ \finite F\ card_insert have "n \ 0" by auto then have "\m. n = Suc m" by presburger then obtain m where n: "n = Suc m" by blast from pFs card_insert have *: "card ?xF = fact n" unfolding xfgpF' card_image[OF ginj] using \finite F\ \finite ?pF\ by (simp only: Collect_case_prod Collect_mem_eq card_cartesian_product) (simp add: n) from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" by (simp add: xfgpF' n) from * have "card ?xF = fact n" unfolding xFf by blast } with insert show ?case by simp qed lemma finite_permutations: assumes "finite S" shows "finite {p. p permutes S}" using card_permutations[OF refl assms] by (auto intro: card_ge_0_finite) subsection \Permutations of index set for iterated operations\ lemma (in comm_monoid_set) permute: assumes "p permutes S" shows "F g S = F (g \ p) S" proof - from \p permutes S\ have "inj p" by (rule permutes_inj) then have "inj_on p S" by (auto intro: subset_inj_on) then have "F g (p ` S) = F (g \ p) S" by (rule reindex) moreover from \p permutes S\ have "p ` S = S" by (rule permutes_image) ultimately show ?thesis by simp qed -subsection \Various combinations of transpositions with 2, 1 and 0 common elements\ - -lemma swap_id_common:" a \ c \ b \ c \ - Fun.swap a b id \ Fun.swap a c id = Fun.swap b c id \ Fun.swap a b id" - by (simp add: fun_eq_iff Fun.swap_def) - -lemma swap_id_common': "a \ b \ a \ c \ - Fun.swap a c id \ Fun.swap b c id = Fun.swap b c id \ Fun.swap a b id" - by (simp add: fun_eq_iff Fun.swap_def) - -lemma swap_id_independent: "a \ c \ a \ d \ b \ c \ b \ d \ - Fun.swap a b id \ Fun.swap c d id = Fun.swap c d id \ Fun.swap a b id" - by (simp add: fun_eq_iff Fun.swap_def) - - subsection \Permutations as transposition sequences\ inductive swapidseq :: "nat \ ('a \ 'a) \ bool" where id[simp]: "swapidseq 0 id" | comp_Suc: "swapidseq n p \ a \ b \ swapidseq (Suc n) (Fun.swap a b id \ p)" declare id[unfolded id_def, simp] definition "permutation p \ (\n. swapidseq n p)" subsection \Some closure properties of the set of permutations, with lengths\ lemma permutation_id[simp]: "permutation id" unfolding permutation_def by (rule exI[where x=0]) simp declare permutation_id[unfolded id_def, simp] lemma swapidseq_swap: "swapidseq (if a = b then 0 else 1) (Fun.swap a b id)" apply clarsimp using comp_Suc[of 0 id a b] apply simp done lemma permutation_swap_id: "permutation (Fun.swap a b id)" proof (cases "a = b") case True then show ?thesis by simp next case False then show ?thesis unfolding permutation_def using swapidseq_swap[of a b] by blast qed lemma swapidseq_comp_add: "swapidseq n p \ swapidseq m q \ swapidseq (n + m) (p \ q)" proof (induct n p arbitrary: m q rule: swapidseq.induct) case (id m q) then show ?case by simp next case (comp_Suc n p a b m q) have eq: "Suc n + m = Suc (n + m)" by arith show ?case apply (simp only: eq comp_assoc) apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems] comp_Suc.hyps(3) apply blast+ done qed lemma permutation_compose: "permutation p \ permutation q \ permutation (p \ q)" unfolding permutation_def using swapidseq_comp_add[of _ p _ q] by metis lemma swapidseq_endswap: "swapidseq n p \ a \ b \ swapidseq (Suc n) (p \ Fun.swap a b id)" by (induct n p rule: swapidseq.induct) (use swapidseq_swap[of a b] in \auto simp add: comp_assoc intro: swapidseq.comp_Suc\) lemma swapidseq_inverse_exists: "swapidseq n p \ \q. swapidseq n q \ p \ q = id \ q \ p = id" proof (induct n p rule: swapidseq.induct) case id then show ?case by (rule exI[where x=id]) simp next case (comp_Suc n p a b) from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast let ?q = "q \ Fun.swap a b id" note H = comp_Suc.hyps from swapidseq_swap[of a b] H(3) have *: "swapidseq 1 (Fun.swap a b id)" by simp from swapidseq_comp_add[OF q(1) *] have **: "swapidseq (Suc n) ?q" by simp have "Fun.swap a b id \ p \ ?q = Fun.swap a b id \ (p \ q) \ Fun.swap a b id" by (simp add: o_assoc) also have "\ = id" by (simp add: q(2)) finally have ***: "Fun.swap a b id \ p \ ?q = id" . have "?q \ (Fun.swap a b id \ p) = q \ (Fun.swap a b id \ Fun.swap a b id) \ p" by (simp only: o_assoc) then have "?q \ (Fun.swap a b id \ p) = id" by (simp add: q(3)) with ** *** show ?case by blast qed lemma swapidseq_inverse: assumes "swapidseq n p" shows "swapidseq n (inv p)" using swapidseq_inverse_exists[OF assms] inv_unique_comp[of p] by auto lemma permutation_inverse: "permutation p \ permutation (inv p)" using permutation_def swapidseq_inverse by blast + +subsection \Various combinations of transpositions with 2, 1 and 0 common elements\ + +lemma swap_id_common:" a \ c \ b \ c \ + Fun.swap a b id \ Fun.swap a c id = Fun.swap b c id \ Fun.swap a b id" + by (simp add: fun_eq_iff Fun.swap_def) + +lemma swap_id_common': "a \ b \ a \ c \ + Fun.swap a c id \ Fun.swap b c id = Fun.swap b c id \ Fun.swap a b id" + by (simp add: fun_eq_iff Fun.swap_def) + +lemma swap_id_independent: "a \ c \ a \ d \ b \ c \ b \ d \ + Fun.swap a b id \ Fun.swap c d id = Fun.swap c d id \ Fun.swap a b id" + by (simp add: fun_eq_iff Fun.swap_def) + + subsection \The identity map only has even transposition sequences\ lemma symmetry_lemma: assumes "\a b c d. P a b c d \ P a b d c" and "\a b c d. a \ b \ c \ d \ a = c \ b = d \ a = c \ b \ d \ a \ c \ b = d \ a \ c \ a \ d \ b \ c \ b \ d \ P a b c d" shows "\a b c d. a \ b \ c \ d \ P a b c d" using assms by metis lemma swap_general: "a \ b \ c \ d \ Fun.swap a b id \ Fun.swap c d id = id \ (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id)" proof - assume neq: "a \ b" "c \ d" have "a \ b \ c \ d \ (Fun.swap a b id \ Fun.swap c d id = id \ (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id))" apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d]) apply (simp_all only: swap_commute) apply (case_tac "a = c \ b = d") apply (clarsimp simp only: swap_commute swap_id_idempotent) apply (case_tac "a = c \ b \ d") apply (rule disjI2) apply (rule_tac x="b" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="b" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) apply (case_tac "a \ c \ b = d") apply (rule disjI2) apply (rule_tac x="c" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="c" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) apply (rule disjI2) apply (rule_tac x="c" in exI) apply (rule_tac x="d" in exI) apply (rule_tac x="b" in exI) apply (clarsimp simp add: fun_eq_iff Fun.swap_def) done with neq show ?thesis by metis qed lemma swapidseq_id_iff[simp]: "swapidseq 0 p \ p = id" using swapidseq.cases[of 0 p "p = id"] by auto lemma swapidseq_cases: "swapidseq n p \ n = 0 \ p = id \ (\a b q m. n = Suc m \ p = Fun.swap a b id \ q \ swapidseq m q \ a \ b)" apply (rule iffI) apply (erule swapidseq.cases[of n p]) apply simp apply (rule disjI2) apply (rule_tac x= "a" in exI) apply (rule_tac x= "b" in exI) apply (rule_tac x= "pa" in exI) apply (rule_tac x= "na" in exI) apply simp apply auto apply (rule comp_Suc, simp_all) done lemma fixing_swapidseq_decrease: assumes "swapidseq n p" and "a \ b" and "(Fun.swap a b id \ p) a = a" shows "n \ 0 \ swapidseq (n - 1) (Fun.swap a b id \ p)" using assms proof (induct n arbitrary: p a b) case 0 then show ?case by (auto simp add: Fun.swap_def fun_upd_def) next case (Suc n p a b) from Suc.prems(1) swapidseq_cases[of "Suc n" p] obtain c d q m where cdqm: "Suc n = Suc m" "p = Fun.swap c d id \ q" "swapidseq m q" "c \ d" "n = m" by auto consider "Fun.swap a b id \ Fun.swap c d id = id" | x y z where "x \ a" "y \ a" "z \ a" "x \ y" "Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id" using swap_general[OF Suc.prems(2) cdqm(4)] by metis then show ?case proof cases case 1 then show ?thesis by (simp only: cdqm o_assoc) (simp add: cdqm) next case prems: 2 then have az: "a \ z" by simp from prems have *: "(Fun.swap x y id \ h) a = a \ h a = a" for h by (simp add: Fun.swap_def) from cdqm(2) have "Fun.swap a b id \ p = Fun.swap a b id \ (Fun.swap c d id \ q)" by simp then have "Fun.swap a b id \ p = Fun.swap x y id \ (Fun.swap a z id \ q)" by (simp add: o_assoc prems) then have "(Fun.swap a b id \ p) a = (Fun.swap x y id \ (Fun.swap a z id \ q)) a" by simp then have "(Fun.swap x y id \ (Fun.swap a z id \ q)) a = a" unfolding Suc by metis then have "(Fun.swap a z id \ q) a = a" by (simp only: *) from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az this] have **: "swapidseq (n - 1) (Fun.swap a z id \ q)" "n \ 0" by blast+ from \n \ 0\ have ***: "Suc n - 1 = Suc (n - 1)" by auto show ?thesis apply (simp only: cdqm(2) prems o_assoc ***) apply (simp only: Suc_not_Zero simp_thms comp_assoc) apply (rule comp_Suc) using ** prems apply blast+ done qed qed lemma swapidseq_identity_even: assumes "swapidseq n (id :: 'a \ 'a)" shows "even n" using \swapidseq n id\ proof (induct n rule: nat_less_induct) case H: (1 n) consider "n = 0" | a b :: 'a and q m where "n = Suc m" "id = Fun.swap a b id \ q" "swapidseq m q" "a \ b" using H(2)[unfolded swapidseq_cases[of n id]] by auto then show ?case proof cases case 1 then show ?thesis by presburger next case h: 2 from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]] have m: "m \ 0" "swapidseq (m - 1) (id :: 'a \ 'a)" by auto from h m have mn: "m - 1 < n" by arith from H(1)[rule_format, OF mn m(2)] h(1) m(1) show ?thesis by presburger qed qed subsection \Therefore we have a welldefined notion of parity\ definition "evenperm p = even (SOME n. swapidseq n p)" lemma swapidseq_even_even: assumes m: "swapidseq m p" and n: "swapidseq n p" shows "even m \ even n" proof - from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]] show ?thesis by arith qed lemma evenperm_unique: assumes p: "swapidseq n p" and n:"even n = b" shows "evenperm p = b" unfolding n[symmetric] evenperm_def apply (rule swapidseq_even_even[where p = p]) apply (rule someI[where x = n]) using p apply blast+ done subsection \And it has the expected composition properties\ lemma evenperm_id[simp]: "evenperm id = True" by (rule evenperm_unique[where n = 0]) simp_all lemma evenperm_swap: "evenperm (Fun.swap a b id) = (a = b)" by (rule evenperm_unique[where n="if a = b then 0 else 1"]) (simp_all add: swapidseq_swap) lemma evenperm_comp: assumes "permutation p" "permutation q" shows "evenperm (p \ q) \ evenperm p = evenperm q" proof - from assms obtain n m where n: "swapidseq n p" and m: "swapidseq m q" unfolding permutation_def by blast have "even (n + m) \ (even n \ even m)" by arith from evenperm_unique[OF n refl] evenperm_unique[OF m refl] and evenperm_unique[OF swapidseq_comp_add[OF n m] this] show ?thesis by blast qed lemma evenperm_inv: assumes "permutation p" shows "evenperm (inv p) = evenperm p" proof - from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast show ?thesis by (rule evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]]) qed subsection \A more abstract characterization of permutations\ -lemma bij_iff: "bij f \ (\x. \!y. f y = x)" - unfolding bij_def inj_def surj_def - apply auto - apply metis - apply metis - done - lemma permutation_bijective: assumes "permutation p" shows "bij p" proof - from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast then show ?thesis unfolding bij_iff apply (auto simp add: fun_eq_iff) apply metis done qed lemma permutation_finite_support: assumes "permutation p" shows "finite {x. p x \ x}" proof - from assms obtain n where "swapidseq n p" unfolding permutation_def by blast then show ?thesis proof (induct n p rule: swapidseq.induct) case id then show ?case by simp next case (comp_Suc n p a b) let ?S = "insert a (insert b {x. p x \ x})" from comp_Suc.hyps(2) have *: "finite ?S" by simp from \a \ b\ have **: "{x. (Fun.swap a b id \ p) x \ x} \ ?S" by (auto simp: Fun.swap_def) show ?case by (rule finite_subset[OF ** *]) qed qed lemma permutation_lemma: assumes "finite S" and "bij p" - and "\x. x\ S \ p x = x" + and "\x. x \ S \ p x = x" shows "permutation p" using assms proof (induct S arbitrary: p rule: finite_induct) case empty then show ?case by simp next case (insert a F p) let ?r = "Fun.swap a (p a) id \ p" let ?q = "Fun.swap a (p a) id \ ?r" have *: "?r a = a" by (simp add: Fun.swap_def) from insert * have **: "\x. x \ F \ ?r x = x" by (metis bij_pointE comp_apply id_apply insert_iff swap_apply(3)) have "bij ?r" by (rule bij_swap_compose_bij[OF insert(4)]) have "permutation ?r" by (rule insert(3)[OF \bij ?r\ **]) then have "permutation ?q" by (simp add: permutation_compose permutation_swap_id) then show ?case by (simp add: o_assoc) qed lemma permutation: "permutation p \ bij p \ finite {x. p x \ x}" (is "?lhs \ ?b \ ?f") proof assume ?lhs with permutation_bijective permutation_finite_support show "?b \ ?f" by auto next assume "?b \ ?f" then have "?f" "?b" by blast+ from permutation_lemma[OF this] show ?lhs by blast qed lemma permutation_inverse_works: assumes "permutation p" shows "inv p \ p = id" and "p \ inv p = id" using permutation_bijective [OF assms] by (auto simp: bij_def inj_iff surj_iff) lemma permutation_inverse_compose: assumes p: "permutation p" and q: "permutation q" shows "inv (p \ q) = inv q \ inv p" proof - note ps = permutation_inverse_works[OF p] note qs = permutation_inverse_works[OF q] have "p \ q \ (inv q \ inv p) = p \ (q \ inv q) \ inv p" by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) finally have *: "p \ q \ (inv q \ inv p) = id" . have "inv q \ inv p \ (p \ q) = inv q \ (inv p \ p) \ q" by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) finally have **: "inv q \ inv p \ (p \ q) = id" . show ?thesis by (rule inv_unique_comp[OF * **]) qed subsection \Relation to \permutes\\ lemma permutation_permutes: "permutation p \ (\S. finite S \ p permutes S)" unfolding permutation permutes_def bij_iff[symmetric] apply (rule iffI, clarify) apply (rule exI[where x="{x. p x \ x}"]) apply simp apply clarsimp apply (rule_tac B="S" in finite_subset) apply auto done subsection \Hence a sort of induction principle composing by swaps\ -lemma permutes_induct: "finite S \ P id \ - (\a b p. a \ S \ b \ S \ P p \ P p \ permutation p \ P (Fun.swap a b id \ p)) \ - (\p. p permutes S \ P p)" -proof (induct S rule: finite_induct) +lemma permutes_induct [consumes 2, case_names id swap]: + \P p\ if \p permutes S\ \finite S\ + and id: \P id\ + and swap: \\a b p. a \ S \ b \ S \ permutation p \ P p \ P (Fun.swap a b id \ p)\ +using \finite S\ \p permutes S\ id swap proof (induction S arbitrary: p) case empty then show ?case by auto next case (insert x F p) let ?r = "Fun.swap x (p x) id \ p" let ?q = "Fun.swap x (p x) id \ ?r" have qp: "?q = p" by (simp add: o_assoc) - from permutes_insert_lemma[OF insert.prems(3)] insert have Pr: "P ?r" + from permutes_insert_lemma[OF \p permutes insert x F\] insert have Pr: "P ?r" by blast - from permutes_in_image[OF insert.prems(3), of x] + from permutes_in_image[OF \p permutes insert x F\, of x] have pxF: "p x \ insert x F" by simp have xF: "x \ insert x F" by simp have rp: "permutation ?r" unfolding permutation_permutes - using insert.hyps(1) permutes_insert_lemma[OF insert.prems(3)] + using insert.hyps(1) permutes_insert_lemma[OF \p permutes insert x F\] by blast - from insert.prems(2)[OF xF pxF Pr Pr rp] qp show ?case + from insert.prems(3)[OF xF pxF rp Pr] qp show ?case by (simp only:) qed subsection \Sign of a permutation as a real number\ -definition "sign p = (if evenperm p then (1::int) else -1)" +definition sign :: \('a \ 'a) \ int\ \ \TODO: prefer less generic name\ + where \sign p = (if evenperm p then (1::int) else -1)\ lemma sign_nz: "sign p \ 0" by (simp add: sign_def) lemma sign_id: "sign id = 1" by (simp add: sign_def) lemma sign_inverse: "permutation p \ sign (inv p) = sign p" by (simp add: sign_def evenperm_inv) lemma sign_compose: "permutation p \ permutation q \ sign (p \ q) = sign p * sign q" by (simp add: sign_def evenperm_comp) lemma sign_swap_id: "sign (Fun.swap a b id) = (if a = b then 1 else -1)" by (simp add: sign_def evenperm_swap) lemma sign_idempotent: "sign p * sign p = 1" by (simp add: sign_def) subsection \Permuting a list\ text \This function permutes a list by applying a permutation to the indices.\ definition permute_list :: "(nat \ nat) \ 'a list \ 'a list" where "permute_list f xs = map (\i. xs ! (f i)) [0.. g) xs = permute_list g (permute_list f xs)" using assms[THEN permutes_in_image] by (auto simp add: permute_list_def) lemma permute_list_ident [simp]: "permute_list (\x. x) xs = xs" by (simp add: permute_list_def map_nth) lemma permute_list_id [simp]: "permute_list id xs = xs" by (simp add: id_def) lemma mset_permute_list [simp]: fixes xs :: "'a list" assumes "f permutes {.. x < length xs" for x using permutes_in_image[OF assms] by auto have "count (mset (permute_list f xs)) y = card ((\i. xs ! f i) -` {y} \ {..i. xs ! f i) -` {y} \ {.. y = xs ! i}" by auto also from assms have "card \ = card {i. i < length xs \ y = xs ! i}" by (intro card_vimage_inj) (auto simp: permutes_inj permutes_surj) also have "\ = count (mset xs) y" by (simp add: count_mset length_filter_conv_card) finally show "count (mset (permute_list f xs)) y = count (mset xs) y" by simp qed lemma set_permute_list [simp]: assumes "f permutes {.. i < length ys" for i by simp have "permute_list f (zip xs ys) = map (\i. zip xs ys ! f i) [0.. = map (\(x, y). (xs ! f x, ys ! f y)) (zip [0.. = zip (permute_list f xs) (permute_list f ys)" by (simp_all add: permute_list_def zip_map_map) finally show ?thesis . qed lemma map_of_permute: assumes "\ permutes fst ` set xs" shows "map_of xs \ \ = map_of (map (\(x,y). (inv \ x, y)) xs)" (is "_ = map_of (map ?f _)") proof from assms have "inj \" "surj \" by (simp_all add: permutes_inj permutes_surj) then show "(map_of xs \ \) x = map_of (map ?f xs) x" for x by (induct xs) (auto simp: inv_f_f surj_f_inv_f) qed subsection \More lemmas about permutations\ text \The following few lemmas were contributed by Lukas Bulwahn.\ lemma count_image_mset_eq_card_vimage: assumes "finite A" shows "count (image_mset f (mset_set A)) b = card {a \ A. f a = b}" using assms proof (induct A) case empty show ?case by simp next case (insert x F) show ?case proof (cases "f x = b") case True with insert.hyps have "count (image_mset f (mset_set (insert x F))) b = Suc (card {a \ F. f a = f x})" by auto also from insert.hyps(1,2) have "\ = card (insert x {a \ F. f a = f x})" by simp also from \f x = b\ have "card (insert x {a \ F. f a = f x}) = card {a \ insert x F. f a = b}" by (auto intro: arg_cong[where f="card"]) finally show ?thesis using insert by auto next case False then have "{a \ F. f a = b} = {a \ insert x F. f a = b}" by auto with insert False show ?thesis by simp qed qed \ \Prove \image_mset_eq_implies_permutes\ ...\ lemma image_mset_eq_implies_permutes: fixes f :: "'a \ 'b" assumes "finite A" and mset_eq: "image_mset f (mset_set A) = image_mset f' (mset_set A)" obtains p where "p permutes A" and "\x\A. f x = f' (p x)" proof - from \finite A\ have [simp]: "finite {a \ A. f a = (b::'b)}" for f b by auto have "f ` A = f' ` A" proof - from \finite A\ have "f ` A = f ` (set_mset (mset_set A))" by simp also have "\ = f' ` set_mset (mset_set A)" by (metis mset_eq multiset.set_map) also from \finite A\ have "\ = f' ` A" by simp finally show ?thesis . qed have "\b\(f ` A). \p. bij_betw p {a \ A. f a = b} {a \ A. f' a = b}" proof fix b from mset_eq have "count (image_mset f (mset_set A)) b = count (image_mset f' (mset_set A)) b" by simp with \finite A\ have "card {a \ A. f a = b} = card {a \ A. f' a = b}" by (simp add: count_image_mset_eq_card_vimage) then show "\p. bij_betw p {a\A. f a = b} {a \ A. f' a = b}" by (intro finite_same_card_bij) simp_all qed then have "\p. \b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" by (rule bchoice) then obtain p where p: "\b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" .. define p' where "p' = (\a. if a \ A then p (f a) a else a)" have "p' permutes A" proof (rule bij_imp_permutes) have "disjoint_family_on (\i. {a \ A. f' a = i}) (f ` A)" by (auto simp: disjoint_family_on_def) moreover have "bij_betw (\a. p (f a) a) {a \ A. f a = b} {a \ A. f' a = b}" if "b \ f ` A" for b using p that by (subst bij_betw_cong[where g="p b"]) auto ultimately have "bij_betw (\a. p (f a) a) (\b\f ` A. {a \ A. f a = b}) (\b\f ` A. {a \ A. f' a = b})" by (rule bij_betw_UNION_disjoint) moreover have "(\b\f ` A. {a \ A. f a = b}) = A" by auto moreover from \f ` A = f' ` A\ have "(\b\f ` A. {a \ A. f' a = b}) = A" by auto ultimately show "bij_betw p' A A" unfolding p'_def by (subst bij_betw_cong[where g="(\a. p (f a) a)"]) auto next show "\x. x \ A \ p' x = x" by (simp add: p'_def) qed moreover from p have "\x\A. f x = f' (p' x)" unfolding p'_def using bij_betwE by fastforce ultimately show ?thesis .. qed lemma mset_set_upto_eq_mset_upto: "mset_set {.. \... and derive the existing property:\ lemma mset_eq_permutation: fixes xs ys :: "'a list" assumes mset_eq: "mset xs = mset ys" obtains p where "p permutes {..i. xs ! i) (mset_set {..i. ys ! i) (mset_set {..i\{..i \ S. p i \ i" shows "p = id" proof - have "p n = n" for n using assms proof (induct n arbitrary: S rule: less_induct) case (less n) show ?case proof (cases "n \ S") case False with less(2) show ?thesis unfolding permutes_def by metis next case True with less(3) have "p n < n \ p n = n" by auto then show ?thesis proof assume "p n < n" with less have "p (p n) = p n" by metis with permutes_inj[OF less(2)] have "p n = n" unfolding inj_def by blast with \p n < n\ have False by simp then show ?thesis .. qed qed qed then show ?thesis by (auto simp: fun_eq_iff) qed lemma permutes_natset_ge: fixes S :: "'a::wellorder set" assumes p: "p permutes S" and le: "\i \ S. p i \ i" shows "p = id" proof - have "i \ inv p i" if "i \ S" for i proof - from that permutes_in_image[OF permutes_inv[OF p]] have "inv p i \ S" by simp with le have "p (inv p i) \ inv p i" by blast with permutes_inverses[OF p] show ?thesis by simp qed then have "\i\S. inv p i \ i" by blast from permutes_natset_le[OF permutes_inv[OF p] this] have "inv p = inv id" by simp then show ?thesis apply (subst permutes_inv_inv[OF p, symmetric]) apply (rule inv_unique_comp) apply simp_all done qed lemma image_inverse_permutations: "{inv p |p. p permutes S} = {p. p permutes S}" apply (rule set_eqI) apply auto using permutes_inv_inv permutes_inv apply auto apply (rule_tac x="inv x" in exI) apply auto done lemma image_compose_permutations_left: assumes "q permutes S" shows "{q \ p |p. p permutes S} = {p. p permutes S}" apply (rule set_eqI) apply auto apply (rule permutes_compose) using assms apply auto apply (rule_tac x = "inv q \ x" in exI) apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o) done lemma image_compose_permutations_right: assumes "q permutes S" shows "{p \ q | p. p permutes S} = {p . p permutes S}" apply (rule set_eqI) apply auto apply (rule permutes_compose) using assms apply auto apply (rule_tac x = "x \ inv q" in exI) apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o comp_assoc) done lemma permutes_in_seg: "p permutes {1 ..n} \ i \ {1..n} \ 1 \ p i \ p i \ n" by (simp add: permutes_def) metis lemma sum_permutations_inverse: "sum f {p. p permutes S} = sum (\p. f(inv p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p . p permutes S}" have *: "inj_on inv ?S" proof (auto simp add: inj_on_def) fix q r assume q: "q permutes S" and r: "r permutes S" and qr: "inv q = inv r" then have "inv (inv q) = inv (inv r)" by simp with permutes_inv_inv[OF q] permutes_inv_inv[OF r] show "q = r" by metis qed have **: "inv ` ?S = ?S" using image_inverse_permutations by blast have ***: "?rhs = sum (f \ inv) ?S" by (simp add: o_def) from sum.reindex[OF *, of f] show ?thesis by (simp only: ** ***) qed lemma setum_permutations_compose_left: assumes q: "q permutes S" shows "sum f {p. p permutes S} = sum (\p. f(q \ p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" have *: "?rhs = sum (f \ ((\) q)) ?S" by (simp add: o_def) have **: "inj_on ((\) q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "q \ p = q \ r" then have "inv q \ q \ p = inv q \ q \ r" by (simp add: comp_assoc) with permutes_inj[OF q, unfolded inj_iff] show "p = r" by simp qed have "((\) q) ` ?S = ?S" using image_compose_permutations_left[OF q] by auto with * sum.reindex[OF **, of f] show ?thesis by (simp only:) qed lemma sum_permutations_compose_right: assumes q: "q permutes S" shows "sum f {p. p permutes S} = sum (\p. f(p \ q)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" have *: "?rhs = sum (f \ (\p. p \ q)) ?S" by (simp add: o_def) have **: "inj_on (\p. p \ q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "p \ q = r \ q" then have "p \ (q \ inv q) = r \ (q \ inv q)" by (simp add: o_assoc) with permutes_surj[OF q, unfolded surj_iff] show "p = r" by simp qed from image_compose_permutations_right[OF q] have "(\p. p \ q) ` ?S = ?S" by auto with * sum.reindex[OF **, of f] show ?thesis by (simp only:) qed subsection \Sum over a set of permutations (could generalize to iteration)\ lemma sum_over_permutations_insert: assumes fS: "finite S" and aS: "a \ S" shows "sum f {p. p permutes (insert a S)} = sum (\b. sum (\q. f (Fun.swap a b id \ q)) {p. p permutes S}) (insert a S)" proof - have *: "\f a b. (\(b, p). f (Fun.swap a b id \ p)) = f \ (\(b,p). Fun.swap a b id \ p)" by (simp add: fun_eq_iff) have **: "\P Q. {(a, b). a \ P \ b \ Q} = P \ Q" by blast show ?thesis unfolding * ** sum.cartesian_product permutes_insert proof (rule sum.reindex) let ?f = "(\(b, y). Fun.swap a b id \ y)" let ?P = "{p. p permutes S}" { fix b c p q assume b: "b \ insert a S" assume c: "c \ insert a S" assume p: "p permutes S" assume q: "q permutes S" assume eq: "Fun.swap a b id \ p = Fun.swap a c id \ q" from p q aS have pa: "p a = a" and qa: "q a = a" unfolding permutes_def by metis+ from eq have "(Fun.swap a b id \ p) a = (Fun.swap a c id \ q) a" by simp then have bc: "b = c" by (simp add: permutes_def pa qa o_def fun_upd_def Fun.swap_def id_def cong del: if_weak_cong split: if_split_asm) from eq[unfolded bc] have "(\p. Fun.swap a c id \ p) (Fun.swap a c id \ p) = (\p. Fun.swap a c id \ p) (Fun.swap a c id \ q)" by simp then have "p = q" unfolding o_assoc swap_id_idempotent by simp with bc have "b = c \ p = q" by blast } then show "inj_on ?f (insert a S \ ?P)" unfolding inj_on_def by clarify metis qed qed subsection \Constructing permutations from association lists\ definition list_permutes :: "('a \ 'a) list \ 'a set \ bool" where "list_permutes xs A \ set (map fst xs) \ A \ set (map snd xs) = set (map fst xs) \ distinct (map fst xs) \ distinct (map snd xs)" lemma list_permutesI [simp]: assumes "set (map fst xs) \ A" "set (map snd xs) = set (map fst xs)" "distinct (map fst xs)" shows "list_permutes xs A" proof - from assms(2,3) have "distinct (map snd xs)" by (intro card_distinct) (simp_all add: distinct_card del: set_map) with assms show ?thesis by (simp add: list_permutes_def) qed definition permutation_of_list :: "('a \ 'a) list \ 'a \ 'a" where "permutation_of_list xs x = (case map_of xs x of None \ x | Some y \ y)" lemma permutation_of_list_Cons: "permutation_of_list ((x, y) # xs) x' = (if x = x' then y else permutation_of_list xs x')" by (simp add: permutation_of_list_def) fun inverse_permutation_of_list :: "('a \ 'a) list \ 'a \ 'a" where "inverse_permutation_of_list [] x = x" | "inverse_permutation_of_list ((y, x') # xs) x = (if x = x' then y else inverse_permutation_of_list xs x)" declare inverse_permutation_of_list.simps [simp del] lemma inj_on_map_of: assumes "distinct (map snd xs)" shows "inj_on (map_of xs) (set (map fst xs))" proof (rule inj_onI) fix x y assume xy: "x \ set (map fst xs)" "y \ set (map fst xs)" assume eq: "map_of xs x = map_of xs y" from xy obtain x' y' where x'y': "map_of xs x = Some x'" "map_of xs y = Some y'" by (cases "map_of xs x"; cases "map_of xs y") (simp_all add: map_of_eq_None_iff) moreover from x'y' have *: "(x, x') \ set xs" "(y, y') \ set xs" by (force dest: map_of_SomeD)+ moreover from * eq x'y' have "x' = y'" by simp ultimately show "x = y" using assms by (force simp: distinct_map dest: inj_onD[of _ _ "(x,x')" "(y,y')"]) qed lemma inj_on_the: "None \ A \ inj_on the A" by (auto simp: inj_on_def option.the_def split: option.splits) lemma inj_on_map_of': assumes "distinct (map snd xs)" shows "inj_on (the \ map_of xs) (set (map fst xs))" by (intro comp_inj_on inj_on_map_of assms inj_on_the) (force simp: eq_commute[of None] map_of_eq_None_iff) lemma image_map_of: assumes "distinct (map fst xs)" shows "map_of xs ` set (map fst xs) = Some ` set (map snd xs)" using assms by (auto simp: rev_image_eqI) lemma the_Some_image [simp]: "the ` Some ` A = A" by (subst image_image) simp lemma image_map_of': assumes "distinct (map fst xs)" shows "(the \ map_of xs) ` set (map fst xs) = set (map snd xs)" by (simp only: image_comp [symmetric] image_map_of assms the_Some_image) lemma permutation_of_list_permutes [simp]: assumes "list_permutes xs A" shows "permutation_of_list xs permutes A" (is "?f permutes _") proof (rule permutes_subset[OF bij_imp_permutes]) from assms show "set (map fst xs) \ A" by (simp add: list_permutes_def) from assms have "inj_on (the \ map_of xs) (set (map fst xs))" (is ?P) by (intro inj_on_map_of') (simp_all add: list_permutes_def) also have "?P \ inj_on ?f (set (map fst xs))" by (intro inj_on_cong) (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits) finally have "bij_betw ?f (set (map fst xs)) (?f ` set (map fst xs))" by (rule inj_on_imp_bij_betw) also from assms have "?f ` set (map fst xs) = (the \ map_of xs) ` set (map fst xs)" by (intro image_cong refl) (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits) also from assms have "\ = set (map fst xs)" by (subst image_map_of') (simp_all add: list_permutes_def) finally show "bij_betw ?f (set (map fst xs)) (set (map fst xs))" . qed (force simp: permutation_of_list_def dest!: map_of_SomeD split: option.splits)+ lemma eval_permutation_of_list [simp]: "permutation_of_list [] x = x" "x = x' \ permutation_of_list ((x',y)#xs) x = y" "x \ x' \ permutation_of_list ((x',y')#xs) x = permutation_of_list xs x" by (simp_all add: permutation_of_list_def) lemma eval_inverse_permutation_of_list [simp]: "inverse_permutation_of_list [] x = x" "x = x' \ inverse_permutation_of_list ((y,x')#xs) x = y" "x \ x' \ inverse_permutation_of_list ((y',x')#xs) x = inverse_permutation_of_list xs x" by (simp_all add: inverse_permutation_of_list.simps) lemma permutation_of_list_id: "x \ set (map fst xs) \ permutation_of_list xs x = x" by (induct xs) (auto simp: permutation_of_list_Cons) lemma permutation_of_list_unique': "distinct (map fst xs) \ (x, y) \ set xs \ permutation_of_list xs x = y" by (induct xs) (force simp: permutation_of_list_Cons)+ lemma permutation_of_list_unique: "list_permutes xs A \ (x, y) \ set xs \ permutation_of_list xs x = y" by (intro permutation_of_list_unique') (simp_all add: list_permutes_def) lemma inverse_permutation_of_list_id: "x \ set (map snd xs) \ inverse_permutation_of_list xs x = x" by (induct xs) auto lemma inverse_permutation_of_list_unique': "distinct (map snd xs) \ (x, y) \ set xs \ inverse_permutation_of_list xs y = x" - by (induct xs) (force simp: inverse_permutation_of_list.simps)+ + by (induct xs) (force simp: inverse_permutation_of_list.simps(2))+ lemma inverse_permutation_of_list_unique: "list_permutes xs A \ (x,y) \ set xs \ inverse_permutation_of_list xs y = x" by (intro inverse_permutation_of_list_unique') (simp_all add: list_permutes_def) lemma inverse_permutation_of_list_correct: fixes A :: "'a set" assumes "list_permutes xs A" shows "inverse_permutation_of_list xs = inv (permutation_of_list xs)" proof (rule ext, rule sym, subst permutes_inv_eq) from assms show "permutation_of_list xs permutes A" by simp show "permutation_of_list xs (inverse_permutation_of_list xs x) = x" for x proof (cases "x \ set (map snd xs)") case True then obtain y where "(y, x) \ set xs" by auto with assms show ?thesis by (simp add: inverse_permutation_of_list_unique permutation_of_list_unique) next case False with assms show ?thesis by (auto simp: list_permutes_def inverse_permutation_of_list_id permutation_of_list_id) qed qed end