# HG changeset patch # User immler # Date 1526376823 -7200 # Node ID 2af1f142f8555e838a6c06ea3dc6a2c9a59809c4 # Parent 48262e3a2bded448da6cc1f2fb703c59e7019799 move FuncSet back to HOL-Library (amending 493b818e8e10) diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Algebra/Congruence.thy --- a/src/HOL/Algebra/Congruence.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Algebra/Congruence.thy Tue May 15 11:33:43 2018 +0200 @@ -5,7 +5,8 @@ theory Congruence imports - Main HOL.FuncSet + Main + "HOL-Library.FuncSet" begin section \Objects\ diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Algebra/Group.thy Tue May 15 11:33:43 2018 +0200 @@ -5,7 +5,7 @@ *) theory Group -imports Complete_Lattice HOL.FuncSet +imports Complete_Lattice "HOL-Library.FuncSet" begin section \Monoids and Groups\ diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Analysis/Finite_Cartesian_Product.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Tue May 15 11:33:43 2018 +0200 @@ -10,6 +10,7 @@ L2_Norm "HOL-Library.Numeral_Type" "HOL-Library.Countable_Set" + "HOL-Library.FuncSet" begin subsection \Finite Cartesian products, with indexing and lambdas\ diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Analysis/Sigma_Algebra.thy --- a/src/HOL/Analysis/Sigma_Algebra.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Analysis/Sigma_Algebra.thy Tue May 15 11:33:43 2018 +0200 @@ -11,6 +11,7 @@ imports Complex_Main "HOL-Library.Countable_Set" + "HOL-Library.FuncSet" "HOL-Library.Indicator_Function" "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Disjoint_Sets" diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue May 15 11:33:43 2018 +0200 @@ -10,6 +10,7 @@ imports "HOL-Library.Indicator_Function" "HOL-Library.Countable_Set" + "HOL-Library.FuncSet" Linear_Algebra Norm_Arith begin diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/FuncSet.thy --- a/src/HOL/FuncSet.thy Tue May 15 06:23:12 2018 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,568 +0,0 @@ -(* Title: HOL/FuncSet.thy - Author: Florian Kammueller and Lawrence C Paulson, Lukas Bulwahn -*) - -section \Pi and Function Sets\ - -theory FuncSet - imports Main - abbrevs PiE = "Pi\<^sub>E" - and PIE = "\\<^sub>E" -begin - -definition Pi :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" - where "Pi A B = {f. \x. x \ A \ f x \ B x}" - -definition extensional :: "'a set \ ('a \ 'b) set" - where "extensional A = {f. \x. x \ A \ f x = undefined}" - -definition "restrict" :: "('a \ 'b) \ 'a set \ 'a \ 'b" - where "restrict f A = (\x. if x \ A then f x else undefined)" - -abbreviation funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 60) - where "A \ B \ Pi A (\_. B)" - -syntax - "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\ _\_./ _)" 10) - "_lam" :: "pttrn \ 'a set \ ('a \ 'b) \ ('a \ 'b)" ("(3\_\_./ _)" [0,0,3] 3) -translations - "\ x\A. B" \ "CONST Pi A (\x. B)" - "\x\A. f" \ "CONST restrict (\x. f) A" - -definition "compose" :: "'a set \ ('b \ 'c) \ ('a \ 'b) \ ('a \ 'c)" - where "compose A g f = (\x\A. g (f x))" - - -subsection \Basic Properties of @{term Pi}\ - -lemma Pi_I[intro!]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" - by (simp add: Pi_def) - -lemma Pi_I'[simp]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" - by (simp add:Pi_def) - -lemma funcsetI: "(\x. x \ A \ f x \ B) \ f \ A \ B" - by (simp add: Pi_def) - -lemma Pi_mem: "f \ Pi A B \ x \ A \ f x \ B x" - by (simp add: Pi_def) - -lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" - unfolding Pi_def by auto - -lemma PiE [elim]: "f \ Pi A B \ (f x \ B x \ Q) \ (x \ A \ Q) \ Q" - by (auto simp: Pi_def) - -lemma Pi_cong: "(\w. w \ A \ f w = g w) \ f \ Pi A B \ g \ Pi A B" - by (auto simp: Pi_def) - -lemma funcset_id [simp]: "(\x. x) \ A \ A" - by auto - -lemma funcset_mem: "f \ A \ B \ x \ A \ f x \ B" - by (simp add: Pi_def) - -lemma funcset_image: "f \ A \ B \ f ` A \ B" - by auto - -lemma image_subset_iff_funcset: "F ` A \ B \ F \ A \ B" - by auto - -lemma Pi_eq_empty[simp]: "(\ x \ A. B x) = {} \ (\x\A. B x = {})" - apply (simp add: Pi_def) - apply auto - txt \Converse direction requires Axiom of Choice to exhibit a function - picking an element from each non-empty @{term "B x"}\ - apply (drule_tac x = "\u. SOME y. y \ B u" in spec) - apply auto - apply (cut_tac P = "\y. y \ B x" in some_eq_ex) - apply auto - done - -lemma Pi_empty [simp]: "Pi {} B = UNIV" - by (simp add: Pi_def) - -lemma Pi_Int: "Pi I E \ Pi I F = (\ i\I. E i \ F i)" - by auto - -lemma Pi_UN: - fixes A :: "nat \ 'i \ 'a set" - assumes "finite I" - and mono: "\i n m. i \ I \ n \ m \ A n i \ A m i" - shows "(\n. Pi I (A n)) = (\ i\I. \n. A n i)" -proof (intro set_eqI iffI) - fix f - assume "f \ (\ i\I. \n. A n i)" - then have "\i\I. \n. f i \ A n i" - by auto - from bchoice[OF this] obtain n where n: "f i \ A (n i) i" if "i \ I" for i - by auto - obtain k where k: "n i \ k" if "i \ I" for i - using \finite I\ finite_nat_set_iff_bounded_le[of "n`I"] by auto - have "f \ Pi I (A k)" - proof (intro Pi_I) - fix i - assume "i \ I" - from mono[OF this, of "n i" k] k[OF this] n[OF this] - show "f i \ A k i" by auto - qed - then show "f \ (\n. Pi I (A n))" - by auto -qed auto - -lemma Pi_UNIV [simp]: "A \ UNIV = UNIV" - by (simp add: Pi_def) - -text \Covariance of Pi-sets in their second argument\ -lemma Pi_mono: "(\x. x \ A \ B x \ C x) \ Pi A B \ Pi A C" - by auto - -text \Contravariance of Pi-sets in their first argument\ -lemma Pi_anti_mono: "A' \ A \ Pi A B \ Pi A' B" - by auto - -lemma prod_final: - assumes 1: "fst \ f \ Pi A B" - and 2: "snd \ f \ Pi A C" - shows "f \ (\ z \ A. B z \ C z)" -proof (rule Pi_I) - fix z - assume z: "z \ A" - have "f z = (fst (f z), snd (f z))" - by simp - also have "\ \ B z \ C z" - by (metis SigmaI PiE o_apply 1 2 z) - finally show "f z \ B z \ C z" . -qed - -lemma Pi_split_domain[simp]: "x \ Pi (I \ J) X \ x \ Pi I X \ x \ Pi J X" - by (auto simp: Pi_def) - -lemma Pi_split_insert_domain[simp]: "x \ Pi (insert i I) X \ x \ Pi I X \ x i \ X i" - by (auto simp: Pi_def) - -lemma Pi_cancel_fupd_range[simp]: "i \ I \ x \ Pi I (B(i := b)) \ x \ Pi I B" - by (auto simp: Pi_def) - -lemma Pi_cancel_fupd[simp]: "i \ I \ x(i := a) \ Pi I B \ x \ Pi I B" - by (auto simp: Pi_def) - -lemma Pi_fupd_iff: "i \ I \ f \ Pi I (B(i := A)) \ f \ Pi (I - {i}) B \ f i \ A" - apply auto - apply (drule_tac x=x in Pi_mem) - apply (simp_all split: if_split_asm) - apply (drule_tac x=i in Pi_mem) - apply (auto dest!: Pi_mem) - done - - -subsection \Composition With a Restricted Domain: @{term compose}\ - -lemma funcset_compose: "f \ A \ B \ g \ B \ C \ compose A g f \ A \ C" - by (simp add: Pi_def compose_def restrict_def) - -lemma compose_assoc: - assumes "f \ A \ B" - and "g \ B \ C" - and "h \ C \ D" - shows "compose A h (compose A g f) = compose A (compose B h g) f" - using assms by (simp add: fun_eq_iff Pi_def compose_def restrict_def) - -lemma compose_eq: "x \ A \ compose A g f x = g (f x)" - by (simp add: compose_def restrict_def) - -lemma surj_compose: "f ` A = B \ g ` B = C \ compose A g f ` A = C" - by (auto simp add: image_def compose_eq) - - -subsection \Bounded Abstraction: @{term restrict}\ - -lemma restrict_cong: "I = J \ (\i. i \ J =simp=> f i = g i) \ restrict f I = restrict g J" - by (auto simp: restrict_def fun_eq_iff simp_implies_def) - -lemma restrict_in_funcset: "(\x. x \ A \ f x \ B) \ (\x\A. f x) \ A \ B" - by (simp add: Pi_def restrict_def) - -lemma restrictI[intro!]: "(\x. x \ A \ f x \ B x) \ (\x\A. f x) \ Pi A B" - by (simp add: Pi_def restrict_def) - -lemma restrict_apply[simp]: "(\y\A. f y) x = (if x \ A then f x else undefined)" - by (simp add: restrict_def) - -lemma restrict_apply': "x \ A \ (\y\A. f y) x = f x" - by simp - -lemma restrict_ext: "(\x. x \ A \ f x = g x) \ (\x\A. f x) = (\x\A. g x)" - by (simp add: fun_eq_iff Pi_def restrict_def) - -lemma restrict_UNIV: "restrict f UNIV = f" - by (simp add: restrict_def) - -lemma inj_on_restrict_eq [simp]: "inj_on (restrict f A) A = inj_on f A" - by (simp add: inj_on_def restrict_def) - -lemma Id_compose: "f \ A \ B \ f \ extensional A \ compose A (\y\B. y) f = f" - by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) - -lemma compose_Id: "g \ A \ B \ g \ extensional A \ compose A g (\x\A. x) = g" - by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) - -lemma image_restrict_eq [simp]: "(restrict f A) ` A = f ` A" - by (auto simp add: restrict_def) - -lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \ B)" - unfolding restrict_def by (simp add: fun_eq_iff) - -lemma restrict_fupd[simp]: "i \ I \ restrict (f (i := x)) I = restrict f I" - by (auto simp: restrict_def) - -lemma restrict_upd[simp]: "i \ I \ (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" - by (auto simp: fun_eq_iff) - -lemma restrict_Pi_cancel: "restrict x I \ Pi I A \ x \ Pi I A" - by (auto simp: restrict_def Pi_def) - - -subsection \Bijections Between Sets\ - -text \The definition of @{const bij_betw} is in \Fun.thy\, but most of -the theorems belong here, or need at least @{term Hilbert_Choice}.\ - -lemma bij_betwI: - assumes "f \ A \ B" - and "g \ B \ A" - and g_f: "\x. x\A \ g (f x) = x" - and f_g: "\y. y\B \ f (g y) = y" - shows "bij_betw f A B" - unfolding bij_betw_def -proof - show "inj_on f A" - by (metis g_f inj_on_def) - have "f ` A \ B" - using \f \ A \ B\ by auto - moreover - have "B \ f ` A" - by auto (metis Pi_mem \g \ B \ A\ f_g image_iff) - ultimately show "f ` A = B" - by blast -qed - -lemma bij_betw_imp_funcset: "bij_betw f A B \ f \ A \ B" - by (auto simp add: bij_betw_def) - -lemma inj_on_compose: "bij_betw f A B \ inj_on g B \ inj_on (compose A g f) A" - by (auto simp add: bij_betw_def inj_on_def compose_eq) - -lemma bij_betw_compose: "bij_betw f A B \ bij_betw g B C \ bij_betw (compose A g f) A C" - apply (simp add: bij_betw_def compose_eq inj_on_compose) - apply (auto simp add: compose_def image_def) - done - -lemma bij_betw_restrict_eq [simp]: "bij_betw (restrict f A) A B = bij_betw f A B" - by (simp add: bij_betw_def) - - -subsection \Extensionality\ - -lemma extensional_empty[simp]: "extensional {} = {\x. undefined}" - unfolding extensional_def by auto - -lemma extensional_arb: "f \ extensional A \ x \ A \ f x = undefined" - by (simp add: extensional_def) - -lemma restrict_extensional [simp]: "restrict f A \ extensional A" - by (simp add: restrict_def extensional_def) - -lemma compose_extensional [simp]: "compose A f g \ extensional A" - by (simp add: compose_def) - -lemma extensionalityI: - assumes "f \ extensional A" - and "g \ extensional A" - and "\x. x \ A \ f x = g x" - shows "f = g" - using assms by (force simp add: fun_eq_iff extensional_def) - -lemma extensional_restrict: "f \ extensional A \ restrict f A = f" - by (rule extensionalityI[OF restrict_extensional]) auto - -lemma extensional_subset: "f \ extensional A \ A \ B \ f \ extensional B" - unfolding extensional_def by auto - -lemma inv_into_funcset: "f ` A = B \ (\x\B. inv_into A f x) \ B \ A" - by (unfold inv_into_def) (fast intro: someI2) - -lemma compose_inv_into_id: "bij_betw f A B \ compose A (\y\B. inv_into A f y) f = (\x\A. x)" - apply (simp add: bij_betw_def compose_def) - apply (rule restrict_ext, auto) - done - -lemma compose_id_inv_into: "f ` A = B \ compose B f (\y\B. inv_into A f y) = (\x\B. x)" - apply (simp add: compose_def) - apply (rule restrict_ext) - apply (simp add: f_inv_into_f) - done - -lemma extensional_insert[intro, simp]: - assumes "a \ extensional (insert i I)" - shows "a(i := b) \ extensional (insert i I)" - using assms unfolding extensional_def by auto - -lemma extensional_Int[simp]: "extensional I \ extensional I' = extensional (I \ I')" - unfolding extensional_def by auto - -lemma extensional_UNIV[simp]: "extensional UNIV = UNIV" - by (auto simp: extensional_def) - -lemma restrict_extensional_sub[intro]: "A \ B \ restrict f A \ extensional B" - unfolding restrict_def extensional_def by auto - -lemma extensional_insert_undefined[intro, simp]: - "a \ extensional (insert i I) \ a(i := undefined) \ extensional I" - unfolding extensional_def by auto - -lemma extensional_insert_cancel[intro, simp]: - "a \ extensional I \ a \ extensional (insert i I)" - unfolding extensional_def by auto - - -subsection \Cardinality\ - -lemma card_inj: "f \ A \ B \ inj_on f A \ finite B \ card A \ card B" - by (rule card_inj_on_le) auto - -lemma card_bij: - assumes "f \ A \ B" "inj_on f A" - and "g \ B \ A" "inj_on g B" - and "finite A" "finite B" - shows "card A = card B" - using assms by (blast intro: card_inj order_antisym) - - -subsection \Extensional Function Spaces\ - -definition PiE :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" - where "PiE S T = Pi S T \ extensional S" - -abbreviation "Pi\<^sub>E A B \ PiE A B" - -syntax - "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\\<^sub>E _\_./ _)" 10) -translations - "\\<^sub>E x\A. B" \ "CONST Pi\<^sub>E A (\x. B)" - -abbreviation extensional_funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\\<^sub>E" 60) - where "A \\<^sub>E B \ (\\<^sub>E i\A. B)" - -lemma extensional_funcset_def: "extensional_funcset S T = (S \ T) \ extensional S" - by (simp add: PiE_def) - -lemma PiE_empty_domain[simp]: "Pi\<^sub>E {} T = {\x. undefined}" - unfolding PiE_def by simp - -lemma PiE_UNIV_domain: "Pi\<^sub>E UNIV T = Pi UNIV T" - unfolding PiE_def by simp - -lemma PiE_empty_range[simp]: "i \ I \ F i = {} \ (\\<^sub>E i\I. F i) = {}" - unfolding PiE_def by auto - -lemma PiE_eq_empty_iff: "Pi\<^sub>E I F = {} \ (\i\I. F i = {})" -proof - assume "Pi\<^sub>E I F = {}" - show "\i\I. F i = {}" - proof (rule ccontr) - assume "\ ?thesis" - then have "\i. \y. (i \ I \ y \ F i) \ (i \ I \ y = undefined)" - by auto - from choice[OF this] - obtain f where " \x. (x \ I \ f x \ F x) \ (x \ I \ f x = undefined)" .. - then have "f \ Pi\<^sub>E I F" - by (auto simp: extensional_def PiE_def) - with \Pi\<^sub>E I F = {}\ show False - by auto - qed -qed (auto simp: PiE_def) - -lemma PiE_arb: "f \ Pi\<^sub>E S T \ x \ S \ f x = undefined" - unfolding PiE_def by auto (auto dest!: extensional_arb) - -lemma PiE_mem: "f \ Pi\<^sub>E S T \ x \ S \ f x \ T x" - unfolding PiE_def by auto - -lemma PiE_fun_upd: "y \ T x \ f \ Pi\<^sub>E S T \ f(x := y) \ Pi\<^sub>E (insert x S) T" - unfolding PiE_def extensional_def by auto - -lemma fun_upd_in_PiE: "x \ S \ f \ Pi\<^sub>E (insert x S) T \ f(x := undefined) \ Pi\<^sub>E S T" - unfolding PiE_def extensional_def by auto - -lemma PiE_insert_eq: "Pi\<^sub>E (insert x S) T = (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" -proof - - { - fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" - then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" - by (auto intro!: image_eqI[where x="(f x, f(x := undefined))"] intro: fun_upd_in_PiE PiE_mem) - } - moreover - { - fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" - then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" - by (auto intro!: image_eqI[where x="(f x, f)"] intro: fun_upd_in_PiE PiE_mem simp: insert_absorb) - } - ultimately show ?thesis - by (auto intro: PiE_fun_upd) -qed - -lemma PiE_Int: "Pi\<^sub>E I A \ Pi\<^sub>E I B = Pi\<^sub>E I (\x. A x \ B x)" - by (auto simp: PiE_def) - -lemma PiE_cong: "(\i. i\I \ A i = B i) \ Pi\<^sub>E I A = Pi\<^sub>E I B" - unfolding PiE_def by (auto simp: Pi_cong) - -lemma PiE_E [elim]: - assumes "f \ Pi\<^sub>E A B" - obtains "x \ A" and "f x \ B x" - | "x \ A" and "f x = undefined" - using assms by (auto simp: Pi_def PiE_def extensional_def) - -lemma PiE_I[intro!]: - "(\x. x \ A \ f x \ B x) \ (\x. x \ A \ f x = undefined) \ f \ Pi\<^sub>E A B" - by (simp add: PiE_def extensional_def) - -lemma PiE_mono: "(\x. x \ A \ B x \ C x) \ Pi\<^sub>E A B \ Pi\<^sub>E A C" - by auto - -lemma PiE_iff: "f \ Pi\<^sub>E I X \ (\i\I. f i \ X i) \ f \ extensional I" - by (simp add: PiE_def Pi_iff) - -lemma PiE_restrict[simp]: "f \ Pi\<^sub>E A B \ restrict f A = f" - by (simp add: extensional_restrict PiE_def) - -lemma restrict_PiE[simp]: "restrict f I \ Pi\<^sub>E I S \ f \ Pi I S" - by (auto simp: PiE_iff) - -lemma PiE_eq_subset: - assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" - and eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" - and "i \ I" - shows "F i \ F' i" -proof - fix x - assume "x \ F i" - with ne have "\j. \y. (j \ I \ y \ F j \ (i = j \ x = y)) \ (j \ I \ y = undefined)" - by auto - from choice[OF this] obtain f - where f: " \j. (j \ I \ f j \ F j \ (i = j \ x = f j)) \ (j \ I \ f j = undefined)" .. - then have "f \ Pi\<^sub>E I F" - by (auto simp: extensional_def PiE_def) - then have "f \ Pi\<^sub>E I F'" - using assms by simp - then show "x \ F' i" - using f \i \ I\ by (auto simp: PiE_def) -qed - -lemma PiE_eq_iff_not_empty: - assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" - shows "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i)" -proof (intro iffI ballI) - fix i - assume eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" - assume i: "i \ I" - show "F i = F' i" - using PiE_eq_subset[of I F F', OF ne eq i] - using PiE_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] - by auto -qed (auto simp: PiE_def) - -lemma PiE_eq_iff: - "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i) \ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" -proof (intro iffI disjCI) - assume eq[simp]: "Pi\<^sub>E I F = Pi\<^sub>E I F'" - assume "\ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" - then have "(\i\I. F i \ {}) \ (\i\I. F' i \ {})" - using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by auto - with PiE_eq_iff_not_empty[of I F F'] show "\i\I. F i = F' i" - by auto -next - assume "(\i\I. F i = F' i) \ (\i\I. F i = {}) \ (\i\I. F' i = {})" - then show "Pi\<^sub>E I F = Pi\<^sub>E I F'" - using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by (auto simp: PiE_def) -qed - -lemma extensional_funcset_fun_upd_restricts_rangeI: - "\y \ S. f x \ f y \ f \ (insert x S) \\<^sub>E T \ f(x := undefined) \ S \\<^sub>E (T - {f x})" - unfolding extensional_funcset_def extensional_def - apply auto - apply (case_tac "x = xa") - apply auto - done - -lemma extensional_funcset_fun_upd_extends_rangeI: - assumes "a \ T" "f \ S \\<^sub>E (T - {a})" - shows "f(x := a) \ insert x S \\<^sub>E T" - using assms unfolding extensional_funcset_def extensional_def by auto - - -subsubsection \Injective Extensional Function Spaces\ - -lemma extensional_funcset_fun_upd_inj_onI: - assumes "f \ S \\<^sub>E (T - {a})" - and "inj_on f S" - shows "inj_on (f(x := a)) S" - using assms - unfolding extensional_funcset_def by (auto intro!: inj_on_fun_updI) - -lemma extensional_funcset_extend_domain_inj_on_eq: - assumes "x \ S" - shows "{f. f \ (insert x S) \\<^sub>E T \ inj_on f (insert x S)} = - (\(y, g). g(x:=y)) ` {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" - using assms - apply (auto del: PiE_I PiE_E) - apply (auto intro: extensional_funcset_fun_upd_inj_onI - extensional_funcset_fun_upd_extends_rangeI del: PiE_I PiE_E) - apply (auto simp add: image_iff inj_on_def) - apply (rule_tac x="xa x" in exI) - apply (auto intro: PiE_mem del: PiE_I PiE_E) - apply (rule_tac x="xa(x := undefined)" in exI) - apply (auto intro!: extensional_funcset_fun_upd_restricts_rangeI) - apply (auto dest!: PiE_mem split: if_split_asm) - done - -lemma extensional_funcset_extend_domain_inj_onI: - assumes "x \ S" - shows "inj_on (\(y, g). g(x := y)) {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" - using assms - apply (auto intro!: inj_onI) - apply (metis fun_upd_same) - apply (metis assms PiE_arb fun_upd_triv fun_upd_upd) - done - - -subsubsection \Cardinality\ - -lemma finite_PiE: "finite S \ (\i. i \ S \ finite (T i)) \ finite (\\<^sub>E i \ S. T i)" - by (induct S arbitrary: T rule: finite_induct) (simp_all add: PiE_insert_eq) - -lemma inj_combinator: "x \ S \ inj_on (\(y, g). g(x := y)) (T x \ Pi\<^sub>E S T)" -proof (safe intro!: inj_onI ext) - fix f y g z - assume "x \ S" - assume fg: "f \ Pi\<^sub>E S T" "g \ Pi\<^sub>E S T" - assume "f(x := y) = g(x := z)" - then have *: "\i. (f(x := y)) i = (g(x := z)) i" - unfolding fun_eq_iff by auto - from this[of x] show "y = z" by simp - fix i from *[of i] \x \ S\ fg show "f i = g i" - by (auto split: if_split_asm simp: PiE_def extensional_def) -qed - -lemma card_PiE: "finite S \ card (\\<^sub>E i \ S. T i) = (\ i\S. card (T i))" -proof (induct rule: finite_induct) - case empty - then show ?case by auto -next - case (insert x S) - then show ?case - by (simp add: PiE_insert_eq inj_combinator card_image card_cartesian_product) -qed - -end diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Library/FuncSet.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/FuncSet.thy Tue May 15 11:33:43 2018 +0200 @@ -0,0 +1,568 @@ +(* Title: HOL/FuncSet.thy + Author: Florian Kammueller and Lawrence C Paulson, Lukas Bulwahn +*) + +section \Pi and Function Sets\ + +theory FuncSet + imports Main + abbrevs PiE = "Pi\<^sub>E" + and PIE = "\\<^sub>E" +begin + +definition Pi :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" + where "Pi A B = {f. \x. x \ A \ f x \ B x}" + +definition extensional :: "'a set \ ('a \ 'b) set" + where "extensional A = {f. \x. x \ A \ f x = undefined}" + +definition "restrict" :: "('a \ 'b) \ 'a set \ 'a \ 'b" + where "restrict f A = (\x. if x \ A then f x else undefined)" + +abbreviation funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 60) + where "A \ B \ Pi A (\_. B)" + +syntax + "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\ _\_./ _)" 10) + "_lam" :: "pttrn \ 'a set \ ('a \ 'b) \ ('a \ 'b)" ("(3\_\_./ _)" [0,0,3] 3) +translations + "\ x\A. B" \ "CONST Pi A (\x. B)" + "\x\A. f" \ "CONST restrict (\x. f) A" + +definition "compose" :: "'a set \ ('b \ 'c) \ ('a \ 'b) \ ('a \ 'c)" + where "compose A g f = (\x\A. g (f x))" + + +subsection \Basic Properties of @{term Pi}\ + +lemma Pi_I[intro!]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" + by (simp add: Pi_def) + +lemma Pi_I'[simp]: "(\x. x \ A \ f x \ B x) \ f \ Pi A B" + by (simp add:Pi_def) + +lemma funcsetI: "(\x. x \ A \ f x \ B) \ f \ A \ B" + by (simp add: Pi_def) + +lemma Pi_mem: "f \ Pi A B \ x \ A \ f x \ B x" + by (simp add: Pi_def) + +lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" + unfolding Pi_def by auto + +lemma PiE [elim]: "f \ Pi A B \ (f x \ B x \ Q) \ (x \ A \ Q) \ Q" + by (auto simp: Pi_def) + +lemma Pi_cong: "(\w. w \ A \ f w = g w) \ f \ Pi A B \ g \ Pi A B" + by (auto simp: Pi_def) + +lemma funcset_id [simp]: "(\x. x) \ A \ A" + by auto + +lemma funcset_mem: "f \ A \ B \ x \ A \ f x \ B" + by (simp add: Pi_def) + +lemma funcset_image: "f \ A \ B \ f ` A \ B" + by auto + +lemma image_subset_iff_funcset: "F ` A \ B \ F \ A \ B" + by auto + +lemma Pi_eq_empty[simp]: "(\ x \ A. B x) = {} \ (\x\A. B x = {})" + apply (simp add: Pi_def) + apply auto + txt \Converse direction requires Axiom of Choice to exhibit a function + picking an element from each non-empty @{term "B x"}\ + apply (drule_tac x = "\u. SOME y. y \ B u" in spec) + apply auto + apply (cut_tac P = "\y. y \ B x" in some_eq_ex) + apply auto + done + +lemma Pi_empty [simp]: "Pi {} B = UNIV" + by (simp add: Pi_def) + +lemma Pi_Int: "Pi I E \ Pi I F = (\ i\I. E i \ F i)" + by auto + +lemma Pi_UN: + fixes A :: "nat \ 'i \ 'a set" + assumes "finite I" + and mono: "\i n m. i \ I \ n \ m \ A n i \ A m i" + shows "(\n. Pi I (A n)) = (\ i\I. \n. A n i)" +proof (intro set_eqI iffI) + fix f + assume "f \ (\ i\I. \n. A n i)" + then have "\i\I. \n. f i \ A n i" + by auto + from bchoice[OF this] obtain n where n: "f i \ A (n i) i" if "i \ I" for i + by auto + obtain k where k: "n i \ k" if "i \ I" for i + using \finite I\ finite_nat_set_iff_bounded_le[of "n`I"] by auto + have "f \ Pi I (A k)" + proof (intro Pi_I) + fix i + assume "i \ I" + from mono[OF this, of "n i" k] k[OF this] n[OF this] + show "f i \ A k i" by auto + qed + then show "f \ (\n. Pi I (A n))" + by auto +qed auto + +lemma Pi_UNIV [simp]: "A \ UNIV = UNIV" + by (simp add: Pi_def) + +text \Covariance of Pi-sets in their second argument\ +lemma Pi_mono: "(\x. x \ A \ B x \ C x) \ Pi A B \ Pi A C" + by auto + +text \Contravariance of Pi-sets in their first argument\ +lemma Pi_anti_mono: "A' \ A \ Pi A B \ Pi A' B" + by auto + +lemma prod_final: + assumes 1: "fst \ f \ Pi A B" + and 2: "snd \ f \ Pi A C" + shows "f \ (\ z \ A. B z \ C z)" +proof (rule Pi_I) + fix z + assume z: "z \ A" + have "f z = (fst (f z), snd (f z))" + by simp + also have "\ \ B z \ C z" + by (metis SigmaI PiE o_apply 1 2 z) + finally show "f z \ B z \ C z" . +qed + +lemma Pi_split_domain[simp]: "x \ Pi (I \ J) X \ x \ Pi I X \ x \ Pi J X" + by (auto simp: Pi_def) + +lemma Pi_split_insert_domain[simp]: "x \ Pi (insert i I) X \ x \ Pi I X \ x i \ X i" + by (auto simp: Pi_def) + +lemma Pi_cancel_fupd_range[simp]: "i \ I \ x \ Pi I (B(i := b)) \ x \ Pi I B" + by (auto simp: Pi_def) + +lemma Pi_cancel_fupd[simp]: "i \ I \ x(i := a) \ Pi I B \ x \ Pi I B" + by (auto simp: Pi_def) + +lemma Pi_fupd_iff: "i \ I \ f \ Pi I (B(i := A)) \ f \ Pi (I - {i}) B \ f i \ A" + apply auto + apply (drule_tac x=x in Pi_mem) + apply (simp_all split: if_split_asm) + apply (drule_tac x=i in Pi_mem) + apply (auto dest!: Pi_mem) + done + + +subsection \Composition With a Restricted Domain: @{term compose}\ + +lemma funcset_compose: "f \ A \ B \ g \ B \ C \ compose A g f \ A \ C" + by (simp add: Pi_def compose_def restrict_def) + +lemma compose_assoc: + assumes "f \ A \ B" + and "g \ B \ C" + and "h \ C \ D" + shows "compose A h (compose A g f) = compose A (compose B h g) f" + using assms by (simp add: fun_eq_iff Pi_def compose_def restrict_def) + +lemma compose_eq: "x \ A \ compose A g f x = g (f x)" + by (simp add: compose_def restrict_def) + +lemma surj_compose: "f ` A = B \ g ` B = C \ compose A g f ` A = C" + by (auto simp add: image_def compose_eq) + + +subsection \Bounded Abstraction: @{term restrict}\ + +lemma restrict_cong: "I = J \ (\i. i \ J =simp=> f i = g i) \ restrict f I = restrict g J" + by (auto simp: restrict_def fun_eq_iff simp_implies_def) + +lemma restrict_in_funcset: "(\x. x \ A \ f x \ B) \ (\x\A. f x) \ A \ B" + by (simp add: Pi_def restrict_def) + +lemma restrictI[intro!]: "(\x. x \ A \ f x \ B x) \ (\x\A. f x) \ Pi A B" + by (simp add: Pi_def restrict_def) + +lemma restrict_apply[simp]: "(\y\A. f y) x = (if x \ A then f x else undefined)" + by (simp add: restrict_def) + +lemma restrict_apply': "x \ A \ (\y\A. f y) x = f x" + by simp + +lemma restrict_ext: "(\x. x \ A \ f x = g x) \ (\x\A. f x) = (\x\A. g x)" + by (simp add: fun_eq_iff Pi_def restrict_def) + +lemma restrict_UNIV: "restrict f UNIV = f" + by (simp add: restrict_def) + +lemma inj_on_restrict_eq [simp]: "inj_on (restrict f A) A = inj_on f A" + by (simp add: inj_on_def restrict_def) + +lemma Id_compose: "f \ A \ B \ f \ extensional A \ compose A (\y\B. y) f = f" + by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) + +lemma compose_Id: "g \ A \ B \ g \ extensional A \ compose A g (\x\A. x) = g" + by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) + +lemma image_restrict_eq [simp]: "(restrict f A) ` A = f ` A" + by (auto simp add: restrict_def) + +lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \ B)" + unfolding restrict_def by (simp add: fun_eq_iff) + +lemma restrict_fupd[simp]: "i \ I \ restrict (f (i := x)) I = restrict f I" + by (auto simp: restrict_def) + +lemma restrict_upd[simp]: "i \ I \ (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" + by (auto simp: fun_eq_iff) + +lemma restrict_Pi_cancel: "restrict x I \ Pi I A \ x \ Pi I A" + by (auto simp: restrict_def Pi_def) + + +subsection \Bijections Between Sets\ + +text \The definition of @{const bij_betw} is in \Fun.thy\, but most of +the theorems belong here, or need at least @{term Hilbert_Choice}.\ + +lemma bij_betwI: + assumes "f \ A \ B" + and "g \ B \ A" + and g_f: "\x. x\A \ g (f x) = x" + and f_g: "\y. y\B \ f (g y) = y" + shows "bij_betw f A B" + unfolding bij_betw_def +proof + show "inj_on f A" + by (metis g_f inj_on_def) + have "f ` A \ B" + using \f \ A \ B\ by auto + moreover + have "B \ f ` A" + by auto (metis Pi_mem \g \ B \ A\ f_g image_iff) + ultimately show "f ` A = B" + by blast +qed + +lemma bij_betw_imp_funcset: "bij_betw f A B \ f \ A \ B" + by (auto simp add: bij_betw_def) + +lemma inj_on_compose: "bij_betw f A B \ inj_on g B \ inj_on (compose A g f) A" + by (auto simp add: bij_betw_def inj_on_def compose_eq) + +lemma bij_betw_compose: "bij_betw f A B \ bij_betw g B C \ bij_betw (compose A g f) A C" + apply (simp add: bij_betw_def compose_eq inj_on_compose) + apply (auto simp add: compose_def image_def) + done + +lemma bij_betw_restrict_eq [simp]: "bij_betw (restrict f A) A B = bij_betw f A B" + by (simp add: bij_betw_def) + + +subsection \Extensionality\ + +lemma extensional_empty[simp]: "extensional {} = {\x. undefined}" + unfolding extensional_def by auto + +lemma extensional_arb: "f \ extensional A \ x \ A \ f x = undefined" + by (simp add: extensional_def) + +lemma restrict_extensional [simp]: "restrict f A \ extensional A" + by (simp add: restrict_def extensional_def) + +lemma compose_extensional [simp]: "compose A f g \ extensional A" + by (simp add: compose_def) + +lemma extensionalityI: + assumes "f \ extensional A" + and "g \ extensional A" + and "\x. x \ A \ f x = g x" + shows "f = g" + using assms by (force simp add: fun_eq_iff extensional_def) + +lemma extensional_restrict: "f \ extensional A \ restrict f A = f" + by (rule extensionalityI[OF restrict_extensional]) auto + +lemma extensional_subset: "f \ extensional A \ A \ B \ f \ extensional B" + unfolding extensional_def by auto + +lemma inv_into_funcset: "f ` A = B \ (\x\B. inv_into A f x) \ B \ A" + by (unfold inv_into_def) (fast intro: someI2) + +lemma compose_inv_into_id: "bij_betw f A B \ compose A (\y\B. inv_into A f y) f = (\x\A. x)" + apply (simp add: bij_betw_def compose_def) + apply (rule restrict_ext, auto) + done + +lemma compose_id_inv_into: "f ` A = B \ compose B f (\y\B. inv_into A f y) = (\x\B. x)" + apply (simp add: compose_def) + apply (rule restrict_ext) + apply (simp add: f_inv_into_f) + done + +lemma extensional_insert[intro, simp]: + assumes "a \ extensional (insert i I)" + shows "a(i := b) \ extensional (insert i I)" + using assms unfolding extensional_def by auto + +lemma extensional_Int[simp]: "extensional I \ extensional I' = extensional (I \ I')" + unfolding extensional_def by auto + +lemma extensional_UNIV[simp]: "extensional UNIV = UNIV" + by (auto simp: extensional_def) + +lemma restrict_extensional_sub[intro]: "A \ B \ restrict f A \ extensional B" + unfolding restrict_def extensional_def by auto + +lemma extensional_insert_undefined[intro, simp]: + "a \ extensional (insert i I) \ a(i := undefined) \ extensional I" + unfolding extensional_def by auto + +lemma extensional_insert_cancel[intro, simp]: + "a \ extensional I \ a \ extensional (insert i I)" + unfolding extensional_def by auto + + +subsection \Cardinality\ + +lemma card_inj: "f \ A \ B \ inj_on f A \ finite B \ card A \ card B" + by (rule card_inj_on_le) auto + +lemma card_bij: + assumes "f \ A \ B" "inj_on f A" + and "g \ B \ A" "inj_on g B" + and "finite A" "finite B" + shows "card A = card B" + using assms by (blast intro: card_inj order_antisym) + + +subsection \Extensional Function Spaces\ + +definition PiE :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" + where "PiE S T = Pi S T \ extensional S" + +abbreviation "Pi\<^sub>E A B \ PiE A B" + +syntax + "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\\<^sub>E _\_./ _)" 10) +translations + "\\<^sub>E x\A. B" \ "CONST Pi\<^sub>E A (\x. B)" + +abbreviation extensional_funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\\<^sub>E" 60) + where "A \\<^sub>E B \ (\\<^sub>E i\A. B)" + +lemma extensional_funcset_def: "extensional_funcset S T = (S \ T) \ extensional S" + by (simp add: PiE_def) + +lemma PiE_empty_domain[simp]: "Pi\<^sub>E {} T = {\x. undefined}" + unfolding PiE_def by simp + +lemma PiE_UNIV_domain: "Pi\<^sub>E UNIV T = Pi UNIV T" + unfolding PiE_def by simp + +lemma PiE_empty_range[simp]: "i \ I \ F i = {} \ (\\<^sub>E i\I. F i) = {}" + unfolding PiE_def by auto + +lemma PiE_eq_empty_iff: "Pi\<^sub>E I F = {} \ (\i\I. F i = {})" +proof + assume "Pi\<^sub>E I F = {}" + show "\i\I. F i = {}" + proof (rule ccontr) + assume "\ ?thesis" + then have "\i. \y. (i \ I \ y \ F i) \ (i \ I \ y = undefined)" + by auto + from choice[OF this] + obtain f where " \x. (x \ I \ f x \ F x) \ (x \ I \ f x = undefined)" .. + then have "f \ Pi\<^sub>E I F" + by (auto simp: extensional_def PiE_def) + with \Pi\<^sub>E I F = {}\ show False + by auto + qed +qed (auto simp: PiE_def) + +lemma PiE_arb: "f \ Pi\<^sub>E S T \ x \ S \ f x = undefined" + unfolding PiE_def by auto (auto dest!: extensional_arb) + +lemma PiE_mem: "f \ Pi\<^sub>E S T \ x \ S \ f x \ T x" + unfolding PiE_def by auto + +lemma PiE_fun_upd: "y \ T x \ f \ Pi\<^sub>E S T \ f(x := y) \ Pi\<^sub>E (insert x S) T" + unfolding PiE_def extensional_def by auto + +lemma fun_upd_in_PiE: "x \ S \ f \ Pi\<^sub>E (insert x S) T \ f(x := undefined) \ Pi\<^sub>E S T" + unfolding PiE_def extensional_def by auto + +lemma PiE_insert_eq: "Pi\<^sub>E (insert x S) T = (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" +proof - + { + fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" + then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" + by (auto intro!: image_eqI[where x="(f x, f(x := undefined))"] intro: fun_upd_in_PiE PiE_mem) + } + moreover + { + fix f assume "f \ Pi\<^sub>E (insert x S) T" "x \ S" + then have "f \ (\(y, g). g(x := y)) ` (T x \ Pi\<^sub>E S T)" + by (auto intro!: image_eqI[where x="(f x, f)"] intro: fun_upd_in_PiE PiE_mem simp: insert_absorb) + } + ultimately show ?thesis + by (auto intro: PiE_fun_upd) +qed + +lemma PiE_Int: "Pi\<^sub>E I A \ Pi\<^sub>E I B = Pi\<^sub>E I (\x. A x \ B x)" + by (auto simp: PiE_def) + +lemma PiE_cong: "(\i. i\I \ A i = B i) \ Pi\<^sub>E I A = Pi\<^sub>E I B" + unfolding PiE_def by (auto simp: Pi_cong) + +lemma PiE_E [elim]: + assumes "f \ Pi\<^sub>E A B" + obtains "x \ A" and "f x \ B x" + | "x \ A" and "f x = undefined" + using assms by (auto simp: Pi_def PiE_def extensional_def) + +lemma PiE_I[intro!]: + "(\x. x \ A \ f x \ B x) \ (\x. x \ A \ f x = undefined) \ f \ Pi\<^sub>E A B" + by (simp add: PiE_def extensional_def) + +lemma PiE_mono: "(\x. x \ A \ B x \ C x) \ Pi\<^sub>E A B \ Pi\<^sub>E A C" + by auto + +lemma PiE_iff: "f \ Pi\<^sub>E I X \ (\i\I. f i \ X i) \ f \ extensional I" + by (simp add: PiE_def Pi_iff) + +lemma PiE_restrict[simp]: "f \ Pi\<^sub>E A B \ restrict f A = f" + by (simp add: extensional_restrict PiE_def) + +lemma restrict_PiE[simp]: "restrict f I \ Pi\<^sub>E I S \ f \ Pi I S" + by (auto simp: PiE_iff) + +lemma PiE_eq_subset: + assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" + and eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" + and "i \ I" + shows "F i \ F' i" +proof + fix x + assume "x \ F i" + with ne have "\j. \y. (j \ I \ y \ F j \ (i = j \ x = y)) \ (j \ I \ y = undefined)" + by auto + from choice[OF this] obtain f + where f: " \j. (j \ I \ f j \ F j \ (i = j \ x = f j)) \ (j \ I \ f j = undefined)" .. + then have "f \ Pi\<^sub>E I F" + by (auto simp: extensional_def PiE_def) + then have "f \ Pi\<^sub>E I F'" + using assms by simp + then show "x \ F' i" + using f \i \ I\ by (auto simp: PiE_def) +qed + +lemma PiE_eq_iff_not_empty: + assumes ne: "\i. i \ I \ F i \ {}" "\i. i \ I \ F' i \ {}" + shows "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i)" +proof (intro iffI ballI) + fix i + assume eq: "Pi\<^sub>E I F = Pi\<^sub>E I F'" + assume i: "i \ I" + show "F i = F' i" + using PiE_eq_subset[of I F F', OF ne eq i] + using PiE_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] + by auto +qed (auto simp: PiE_def) + +lemma PiE_eq_iff: + "Pi\<^sub>E I F = Pi\<^sub>E I F' \ (\i\I. F i = F' i) \ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" +proof (intro iffI disjCI) + assume eq[simp]: "Pi\<^sub>E I F = Pi\<^sub>E I F'" + assume "\ ((\i\I. F i = {}) \ (\i\I. F' i = {}))" + then have "(\i\I. F i \ {}) \ (\i\I. F' i \ {})" + using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by auto + with PiE_eq_iff_not_empty[of I F F'] show "\i\I. F i = F' i" + by auto +next + assume "(\i\I. F i = F' i) \ (\i\I. F i = {}) \ (\i\I. F' i = {})" + then show "Pi\<^sub>E I F = Pi\<^sub>E I F'" + using PiE_eq_empty_iff[of I F] PiE_eq_empty_iff[of I F'] by (auto simp: PiE_def) +qed + +lemma extensional_funcset_fun_upd_restricts_rangeI: + "\y \ S. f x \ f y \ f \ (insert x S) \\<^sub>E T \ f(x := undefined) \ S \\<^sub>E (T - {f x})" + unfolding extensional_funcset_def extensional_def + apply auto + apply (case_tac "x = xa") + apply auto + done + +lemma extensional_funcset_fun_upd_extends_rangeI: + assumes "a \ T" "f \ S \\<^sub>E (T - {a})" + shows "f(x := a) \ insert x S \\<^sub>E T" + using assms unfolding extensional_funcset_def extensional_def by auto + + +subsubsection \Injective Extensional Function Spaces\ + +lemma extensional_funcset_fun_upd_inj_onI: + assumes "f \ S \\<^sub>E (T - {a})" + and "inj_on f S" + shows "inj_on (f(x := a)) S" + using assms + unfolding extensional_funcset_def by (auto intro!: inj_on_fun_updI) + +lemma extensional_funcset_extend_domain_inj_on_eq: + assumes "x \ S" + shows "{f. f \ (insert x S) \\<^sub>E T \ inj_on f (insert x S)} = + (\(y, g). g(x:=y)) ` {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" + using assms + apply (auto del: PiE_I PiE_E) + apply (auto intro: extensional_funcset_fun_upd_inj_onI + extensional_funcset_fun_upd_extends_rangeI del: PiE_I PiE_E) + apply (auto simp add: image_iff inj_on_def) + apply (rule_tac x="xa x" in exI) + apply (auto intro: PiE_mem del: PiE_I PiE_E) + apply (rule_tac x="xa(x := undefined)" in exI) + apply (auto intro!: extensional_funcset_fun_upd_restricts_rangeI) + apply (auto dest!: PiE_mem split: if_split_asm) + done + +lemma extensional_funcset_extend_domain_inj_onI: + assumes "x \ S" + shows "inj_on (\(y, g). g(x := y)) {(y, g). y \ T \ g \ S \\<^sub>E (T - {y}) \ inj_on g S}" + using assms + apply (auto intro!: inj_onI) + apply (metis fun_upd_same) + apply (metis assms PiE_arb fun_upd_triv fun_upd_upd) + done + + +subsubsection \Cardinality\ + +lemma finite_PiE: "finite S \ (\i. i \ S \ finite (T i)) \ finite (\\<^sub>E i \ S. T i)" + by (induct S arbitrary: T rule: finite_induct) (simp_all add: PiE_insert_eq) + +lemma inj_combinator: "x \ S \ inj_on (\(y, g). g(x := y)) (T x \ Pi\<^sub>E S T)" +proof (safe intro!: inj_onI ext) + fix f y g z + assume "x \ S" + assume fg: "f \ Pi\<^sub>E S T" "g \ Pi\<^sub>E S T" + assume "f(x := y) = g(x := z)" + then have *: "\i. (f(x := y)) i = (g(x := z)) i" + unfolding fun_eq_iff by auto + from this[of x] show "y = z" by simp + fix i from *[of i] \x \ S\ fg show "f i = g i" + by (auto split: if_split_asm simp: PiE_def extensional_def) +qed + +lemma card_PiE: "finite S \ card (\\<^sub>E i \ S. T i) = (\ i\S. card (T i))" +proof (induct rule: finite_induct) + case empty + then show ?case by auto +next + case (insert x S) + then show ?case + by (simp add: PiE_insert_eq inj_combinator card_image card_cartesian_product) +qed + +end diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Library/Library.thy Tue May 15 11:33:43 2018 +0200 @@ -30,6 +30,7 @@ Finite_Map Float FSet + FuncSet Function_Division Fun_Lexorder Going_To_Filter diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Metis_Examples/Abstraction.thy --- a/src/HOL/Metis_Examples/Abstraction.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Metis_Examples/Abstraction.thy Tue May 15 11:33:43 2018 +0200 @@ -8,7 +8,7 @@ section \Example Featuring Metis's Support for Lambda-Abstractions\ theory Abstraction -imports HOL.FuncSet +imports "HOL-Library.FuncSet" begin (* For Christoph Benzmüller *) diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Metis_Examples/Tarski.thy --- a/src/HOL/Metis_Examples/Tarski.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Metis_Examples/Tarski.thy Tue May 15 11:33:43 2018 +0200 @@ -8,7 +8,7 @@ section \Metis Example Featuring the Full Theorem of Tarski\ theory Tarski -imports Main HOL.FuncSet +imports Main "HOL-Library.FuncSet" begin declare [[metis_new_skolem]] diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Number_Theory/Prime_Powers.thy --- a/src/HOL/Number_Theory/Prime_Powers.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Number_Theory/Prime_Powers.thy Tue May 15 11:33:43 2018 +0200 @@ -6,7 +6,7 @@ *) section \Prime powers\ theory Prime_Powers - imports Complex_Main "HOL-Computational_Algebra.Primes" + imports Complex_Main "HOL-Computational_Algebra.Primes" "HOL-Library.FuncSet" begin definition aprimedivisor :: "'a :: normalization_semidom \ 'a" where diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/Vector_Spaces.thy --- a/src/HOL/Vector_Spaces.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/Vector_Spaces.thy Tue May 15 11:33:43 2018 +0200 @@ -9,7 +9,7 @@ section \Vector Spaces\ theory Vector_Spaces - imports Modules FuncSet + imports Modules begin lemma isomorphism_expand: @@ -847,7 +847,7 @@ lemma linear_exists_left_inverse_on: assumes lf: "linear s1 s2 f" assumes V: "vs1.subspace V" and f: "inj_on f V" - shows "\g\UNIV \ V. linear s2 s1 g \ (\v\V. g (f v) = v)" + shows "\g. g ` UNIV \ V \ linear s2 s1 g \ (\v\V. g (f v) = v)" proof - interpret linear s1 s2 f by fact obtain B where V_eq: "V = vs1.span B" and B: "vs1.independent B" @@ -856,7 +856,7 @@ have f: "inj_on f (vs1.span B)" using f unfolding V_eq . show ?thesis - proof (intro bexI ballI conjI) + proof (intro exI ballI conjI) interpret p: vector_space_pair s2 s1 by unfold_locales have fB: "vs2.independent (f ` B)" using independent_injective_image[OF B f] . @@ -868,7 +868,7 @@ moreover have "the_inv_into B f ` f ` B = B" by (auto simp: image_comp comp_def the_inv_into_f_f inj_on_subset[OF f vs1.span_superset] cong: image_cong) - ultimately show "?g \ UNIV \ V" + ultimately show "?g ` UNIV \ V" by (auto simp: V_eq) have "(?g \ f) v = id v" if "v \ vs1.span B" for v proof (rule vector_space_pair.linear_eq_on[where x=v]) @@ -890,7 +890,7 @@ lemma linear_exists_right_inverse_on: assumes lf: "linear s1 s2 f" assumes "vs1.subspace V" - shows "\g\UNIV \ V. linear s2 s1 g \ (\v\f ` V. f (g v) = v)" + shows "\g. g ` UNIV \ V \ linear s2 s1 g \ (\v\f ` V. f (g v) = v)" proof - obtain B where V_eq: "V = vs1.span B" and B: "vs1.independent B" using vs1.maximal_independent_subset[of V] vs1.span_minimal[OF _ \vs1.subspace V\] @@ -900,7 +900,7 @@ then have "\v\C. \b\B. v = f b" by auto then obtain g where g: "\v. v \ C \ g v \ B" "\v. v \ C \ f (g v) = v" by metis show ?thesis - proof (intro bexI ballI conjI) + proof (intro exI ballI conjI) interpret p: vector_space_pair s2 s1 by unfold_locales let ?g = "p.construct C g" show "linear ( *b) ( *a) ?g" @@ -908,7 +908,7 @@ have "?g v \ vs1.span (g ` C)" for v by (rule p.construct_in_span[OF C]) also have "\ \ V" unfolding V_eq using g by (intro vs1.span_mono) auto - finally show "?g \ UNIV \ V" by auto + finally show "?g ` UNIV \ V" by auto have "(f \ ?g) v = id v" if v: "v \ f ` V" for v proof (rule vector_space_pair.linear_eq_on[where x=v]) show "vector_space_pair ( *b) ( *b)" by unfold_locales @@ -946,7 +946,7 @@ assumes sf: "vs2.span T \ f`vs1.span S" shows "\g. range g \ vs1.span S \ linear s2 s1 g \ (\x\vs2.span T. f (g x) = x)" using linear_exists_right_inverse_on[OF lf vs1.subspace_span, of S] sf - by (auto simp: linear_iff_module_hom) + by (force simp: linear_iff_module_hom) lemma linear_surjective_right_inverse: "linear s1 s2 f \ surj f \ \g. linear s2 s1 g \ f \ g = id" using linear_surj_right_inverse[of f UNIV UNIV] diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/ex/Ballot.thy --- a/src/HOL/ex/Ballot.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/ex/Ballot.thy Tue May 15 11:33:43 2018 +0200 @@ -8,6 +8,7 @@ theory Ballot imports Complex_Main + "HOL-Library.FuncSet" begin subsection \Preliminaries\ diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/ex/Birthday_Paradox.thy --- a/src/HOL/ex/Birthday_Paradox.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/ex/Birthday_Paradox.thy Tue May 15 11:33:43 2018 +0200 @@ -5,7 +5,7 @@ section \A Formulation of the Birthday Paradox\ theory Birthday_Paradox -imports Main HOL.FuncSet +imports Main "HOL-Library.FuncSet" begin section \Cardinality\ diff -r 48262e3a2bde -r 2af1f142f855 src/HOL/ex/Tarski.thy --- a/src/HOL/ex/Tarski.thy Tue May 15 06:23:12 2018 +0200 +++ b/src/HOL/ex/Tarski.thy Tue May 15 11:33:43 2018 +0200 @@ -5,7 +5,7 @@ section \The Full Theorem of Tarski\ theory Tarski -imports Main HOL.FuncSet +imports Main "HOL-Library.FuncSet" begin text \