paulson@13586: (* Title: HOL/Library/FuncSet.thy bulwahn@40631: Author: Florian Kammueller and Lawrence C Paulson, Lukas Bulwahn paulson@13586: *) paulson@13586: wenzelm@14706: header {* Pi and Function Sets *} paulson@13586: nipkow@15131: theory FuncSet haftmann@30663: imports Hilbert_Choice Main nipkow@15131: begin paulson@13586: wenzelm@19736: definition wenzelm@21404: Pi :: "['a set, 'a => 'b set] => ('a => 'b) set" where wenzelm@19736: "Pi A B = {f. \x. x \ A --> f x \ B x}" paulson@13586: wenzelm@21404: definition wenzelm@21404: extensional :: "'a set => ('a => 'b) set" where haftmann@28524: "extensional A = {f. \x. x~:A --> f x = undefined}" paulson@13586: wenzelm@21404: definition wenzelm@21404: "restrict" :: "['a => 'b, 'a set] => ('a => 'b)" where haftmann@28524: "restrict f A = (%x. if x \ A then f x else undefined)" paulson@13586: wenzelm@19536: abbreviation wenzelm@21404: funcset :: "['a set, 'b set] => ('a => 'b) set" wenzelm@21404: (infixr "->" 60) where wenzelm@19536: "A -> B == Pi A (%_. B)" wenzelm@19536: wenzelm@21210: notation (xsymbols) wenzelm@19656: funcset (infixr "\" 60) wenzelm@19536: paulson@13586: syntax wenzelm@19736: "_Pi" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PI _:_./ _)" 10) wenzelm@19736: "_lam" :: "[pttrn, 'a set, 'a => 'b] => ('a=>'b)" ("(3%_:_./ _)" [0,0,3] 3) paulson@13586: paulson@13586: syntax (xsymbols) wenzelm@19736: "_Pi" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\ _\_./ _)" 10) wenzelm@19736: "_lam" :: "[pttrn, 'a set, 'a => 'b] => ('a=>'b)" ("(3\_\_./ _)" [0,0,3] 3) paulson@13586: kleing@14565: syntax (HTML output) wenzelm@19736: "_Pi" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\ _\_./ _)" 10) wenzelm@19736: "_lam" :: "[pttrn, 'a set, 'a => 'b] => ('a=>'b)" ("(3\_\_./ _)" [0,0,3] 3) kleing@14565: paulson@13586: translations wenzelm@20770: "PI x:A. B" == "CONST Pi A (%x. B)" wenzelm@20770: "%x:A. f" == "CONST restrict (%x. f) A" paulson@13586: wenzelm@19736: definition wenzelm@21404: "compose" :: "['a set, 'b => 'c, 'a => 'b] => ('a => 'c)" where wenzelm@19736: "compose A g f = (\x\A. g (f x))" paulson@13586: paulson@13586: paulson@13586: subsection{*Basic Properties of @{term Pi}*} paulson@13586: nipkow@31754: lemma Pi_I[intro!]: "(!!x. x \ A ==> f x \ B x) ==> f \ Pi A B" wenzelm@14706: by (simp add: Pi_def) paulson@13586: nipkow@31731: lemma Pi_I'[simp]: "(!!x. x : A --> f x : B x) ==> f : Pi A B" nipkow@31731: by(simp add:Pi_def) nipkow@31731: paulson@13586: lemma funcsetI: "(!!x. x \ A ==> f x \ B) ==> f \ A -> B" wenzelm@14706: by (simp add: Pi_def) paulson@13586: paulson@13586: lemma Pi_mem: "[|f: Pi A B; x \ A|] ==> f x \ B x" wenzelm@14706: by (simp add: Pi_def) paulson@13586: hoelzl@47761: lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" hoelzl@47761: unfolding Pi_def by auto hoelzl@47761: nipkow@31759: lemma PiE [elim]: nipkow@31754: "f : Pi A B ==> (f x : B x ==> Q) ==> (x ~: A ==> Q) ==> Q" nipkow@31754: by(auto simp: Pi_def) nipkow@31754: hoelzl@38656: lemma Pi_cong: hoelzl@38656: "(\ w. w \ A \ f w = g w) \ f \ Pi A B \ g \ Pi A B" hoelzl@38656: by (auto simp: Pi_def) hoelzl@38656: haftmann@31769: lemma funcset_id [simp]: "(\x. x) \ A \ A" wenzelm@44382: by auto haftmann@31769: paulson@13586: lemma funcset_mem: "[|f \ A -> B; x \ A|] ==> f x \ B" wenzelm@14706: by (simp add: Pi_def) paulson@13586: paulson@14762: lemma funcset_image: "f \ A\B ==> f ` A \ B" nipkow@31754: by auto paulson@14762: nipkow@31754: lemma Pi_eq_empty[simp]: "((PI x: A. B x) = {}) = (\x\A. B(x) = {})" paulson@13593: apply (simp add: Pi_def, auto) paulson@13586: txt{*Converse direction requires Axiom of Choice to exhibit a function paulson@13586: picking an element from each non-empty @{term "B x"}*} paulson@13593: apply (drule_tac x = "%u. SOME y. y \ B u" in spec, auto) wenzelm@14706: apply (cut_tac P= "%y. y \ B x" in some_eq_ex, auto) paulson@13586: done paulson@13586: paulson@13593: lemma Pi_empty [simp]: "Pi {} B = UNIV" nipkow@31754: by (simp add: Pi_def) paulson@13593: paulson@13593: lemma Pi_UNIV [simp]: "A -> UNIV = UNIV" nipkow@31754: by (simp add: Pi_def) nipkow@31727: (* nipkow@31727: lemma funcset_id [simp]: "(%x. x): A -> A" nipkow@31727: by (simp add: Pi_def) nipkow@31727: *) paulson@13586: text{*Covariance of Pi-sets in their second argument*} paulson@13586: lemma Pi_mono: "(!!x. x \ A ==> B x <= C x) ==> Pi A B <= Pi A C" nipkow@31754: by auto paulson@13586: paulson@13586: text{*Contravariance of Pi-sets in their first argument*} paulson@13586: lemma Pi_anti_mono: "A' <= A ==> Pi A B <= Pi A' B" nipkow@31754: by auto paulson@13586: paulson@33271: lemma prod_final: paulson@33271: assumes 1: "fst \ f \ Pi A B" and 2: "snd \ f \ Pi A C" paulson@33271: shows "f \ (\ z \ A. B z \ C z)" paulson@33271: proof (rule Pi_I) paulson@33271: fix z paulson@33271: assume z: "z \ A" paulson@33271: have "f z = (fst (f z), snd (f z))" paulson@33271: by simp paulson@33271: also have "... \ B z \ C z" paulson@33271: by (metis SigmaI PiE o_apply 1 2 z) paulson@33271: finally show "f z \ B z \ C z" . paulson@33271: qed paulson@33271: paulson@13586: paulson@13586: subsection{*Composition With a Restricted Domain: @{term compose}*} paulson@13586: wenzelm@14706: lemma funcset_compose: nipkow@31754: "[| f \ A -> B; g \ B -> C |]==> compose A g f \ A -> C" nipkow@31754: by (simp add: Pi_def compose_def restrict_def) paulson@13586: paulson@13586: lemma compose_assoc: wenzelm@14706: "[| f \ A -> B; g \ B -> C; h \ C -> D |] paulson@13586: ==> compose A h (compose A g f) = compose A (compose B h g) f" nipkow@39302: by (simp add: fun_eq_iff Pi_def compose_def restrict_def) paulson@13586: paulson@13586: lemma compose_eq: "x \ A ==> compose A g f x = g(f(x))" nipkow@31754: by (simp add: compose_def restrict_def) paulson@13586: paulson@13586: lemma surj_compose: "[| f ` A = B; g ` B = C |] ==> compose A g f ` A = C" wenzelm@14706: by (auto simp add: image_def compose_eq) paulson@13586: paulson@13586: paulson@13586: subsection{*Bounded Abstraction: @{term restrict}*} paulson@13586: paulson@13586: lemma restrict_in_funcset: "(!!x. x \ A ==> f x \ B) ==> (\x\A. f x) \ A -> B" wenzelm@14706: by (simp add: Pi_def restrict_def) paulson@13586: nipkow@31754: lemma restrictI[intro!]: "(!!x. x \ A ==> f x \ B x) ==> (\x\A. f x) \ Pi A B" wenzelm@14706: by (simp add: Pi_def restrict_def) paulson@13586: paulson@13586: lemma restrict_apply [simp]: haftmann@28524: "(\y\A. f y) x = (if x \ A then f x else undefined)" wenzelm@14706: by (simp add: restrict_def) paulson@13586: wenzelm@14706: lemma restrict_ext: paulson@13586: "(!!x. x \ A ==> f x = g x) ==> (\x\A. f x) = (\x\A. g x)" nipkow@39302: by (simp add: fun_eq_iff Pi_def restrict_def) paulson@13586: paulson@14853: lemma inj_on_restrict_eq [simp]: "inj_on (restrict f A) A = inj_on f A" wenzelm@14706: by (simp add: inj_on_def restrict_def) paulson@13586: paulson@13586: lemma Id_compose: wenzelm@14706: "[|f \ A -> B; f \ extensional A|] ==> compose A (\y\B. y) f = f" nipkow@39302: by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) paulson@13586: paulson@13586: lemma compose_Id: wenzelm@14706: "[|g \ A -> B; g \ extensional A|] ==> compose A g (\x\A. x) = g" nipkow@39302: by (auto simp add: fun_eq_iff compose_def extensional_def Pi_def) paulson@13586: paulson@14853: lemma image_restrict_eq [simp]: "(restrict f A) ` A = f ` A" wenzelm@19736: by (auto simp add: restrict_def) paulson@13586: paulson@14745: paulson@14762: subsection{*Bijections Between Sets*} paulson@14762: nipkow@26106: text{*The definition of @{const bij_betw} is in @{text "Fun.thy"}, but most of paulson@14762: the theorems belong here, or need at least @{term Hilbert_Choice}.*} paulson@14762: nipkow@39595: lemma bij_betwI: nipkow@39595: assumes "f \ A \ B" and "g \ B \ A" nipkow@39595: and g_f: "\x. x\A \ g (f x) = x" and f_g: "\y. y\B \ f (g y) = y" nipkow@39595: shows "bij_betw f A B" nipkow@39595: unfolding bij_betw_def nipkow@39595: proof nipkow@39595: show "inj_on f A" by (metis g_f inj_on_def) nipkow@39595: next nipkow@39595: have "f ` A \ B" using `f \ A \ B` by auto nipkow@39595: moreover nipkow@39595: have "B \ f ` A" by auto (metis Pi_mem `g \ B \ A` f_g image_iff) nipkow@39595: ultimately show "f ` A = B" by blast nipkow@39595: qed nipkow@39595: paulson@14762: lemma bij_betw_imp_funcset: "bij_betw f A B \ f \ A \ B" nipkow@32988: by (auto simp add: bij_betw_def) paulson@14762: paulson@14853: lemma inj_on_compose: nipkow@31754: "[| bij_betw f A B; inj_on g B |] ==> inj_on (compose A g f) A" nipkow@31754: by (auto simp add: bij_betw_def inj_on_def compose_eq) paulson@14853: paulson@14762: lemma bij_betw_compose: nipkow@31754: "[| bij_betw f A B; bij_betw g B C |] ==> bij_betw (compose A g f) A C" nipkow@31754: apply (simp add: bij_betw_def compose_eq inj_on_compose) nipkow@31754: apply (auto simp add: compose_def image_def) nipkow@31754: done paulson@14762: paulson@14853: lemma bij_betw_restrict_eq [simp]: nipkow@31754: "bij_betw (restrict f A) A B = bij_betw f A B" nipkow@31754: by (simp add: bij_betw_def) paulson@14853: paulson@14853: paulson@14853: subsection{*Extensionality*} paulson@14853: haftmann@28524: lemma extensional_arb: "[|f \ extensional A; x\ A|] ==> f x = undefined" nipkow@31754: by (simp add: extensional_def) paulson@14853: paulson@14853: lemma restrict_extensional [simp]: "restrict f A \ extensional A" nipkow@31754: by (simp add: restrict_def extensional_def) paulson@14853: paulson@14853: lemma compose_extensional [simp]: "compose A f g \ extensional A" nipkow@31754: by (simp add: compose_def) paulson@14853: paulson@14853: lemma extensionalityI: nipkow@31754: "[| f \ extensional A; g \ extensional A; paulson@14853: !!x. x\A ==> f x = g x |] ==> f = g" nipkow@39302: by (force simp add: fun_eq_iff extensional_def) paulson@14853: nipkow@39595: lemma extensional_restrict: "f \ extensional A \ restrict f A = f" nipkow@39595: by(rule extensionalityI[OF restrict_extensional]) auto nipkow@39595: nipkow@33057: lemma inv_into_funcset: "f ` A = B ==> (\x\B. inv_into A f x) : B -> A" nipkow@33057: by (unfold inv_into_def) (fast intro: someI2) paulson@14853: nipkow@33057: lemma compose_inv_into_id: nipkow@33057: "bij_betw f A B ==> compose A (\y\B. inv_into A f y) f = (\x\A. x)" nipkow@31754: apply (simp add: bij_betw_def compose_def) nipkow@31754: apply (rule restrict_ext, auto) nipkow@31754: done paulson@14853: nipkow@33057: lemma compose_id_inv_into: nipkow@33057: "f ` A = B ==> compose B f (\y\B. inv_into A f y) = (\x\B. x)" nipkow@31754: apply (simp add: compose_def) nipkow@31754: apply (rule restrict_ext) nipkow@33057: apply (simp add: f_inv_into_f) nipkow@31754: done paulson@14853: paulson@14762: paulson@14745: subsection{*Cardinality*} paulson@14745: paulson@14745: lemma card_inj: "[|f \ A\B; inj_on f A; finite B|] ==> card(A) \ card(B)" nipkow@31754: by (rule card_inj_on_le) auto paulson@14745: paulson@14745: lemma card_bij: nipkow@31754: "[|f \ A\B; inj_on f A; nipkow@31754: g \ B\A; inj_on g B; finite A; finite B|] ==> card(A) = card(B)" nipkow@31754: by (blast intro: card_inj order_antisym) paulson@14745: bulwahn@40631: subsection {* Extensional Function Spaces *} bulwahn@40631: bulwahn@40631: definition extensional_funcset bulwahn@40631: where "extensional_funcset S T = (S -> T) \ (extensional S)" bulwahn@40631: bulwahn@40631: lemma extensional_empty[simp]: "extensional {} = {%x. undefined}" wenzelm@44382: unfolding extensional_def by auto bulwahn@40631: bulwahn@40631: lemma extensional_funcset_empty_domain: "extensional_funcset {} T = {%x. undefined}" bulwahn@40631: unfolding extensional_funcset_def by simp bulwahn@40631: bulwahn@40631: lemma extensional_funcset_empty_range: bulwahn@40631: assumes "S \ {}" bulwahn@40631: shows "extensional_funcset S {} = {}" bulwahn@40631: using assms unfolding extensional_funcset_def by auto bulwahn@40631: bulwahn@40631: lemma extensional_funcset_arb: bulwahn@40631: assumes "f \ extensional_funcset S T" "x \ S" bulwahn@40631: shows "f x = undefined" bulwahn@40631: using assms bulwahn@40631: unfolding extensional_funcset_def by auto (auto dest!: extensional_arb) bulwahn@40631: bulwahn@40631: lemma extensional_funcset_mem: "f \ extensional_funcset S T \ x \ S \ f x \ T" bulwahn@40631: unfolding extensional_funcset_def by auto bulwahn@40631: bulwahn@40631: lemma extensional_subset: "f : extensional A ==> A <= B ==> f : extensional B" bulwahn@40631: unfolding extensional_def by auto bulwahn@40631: bulwahn@40631: lemma extensional_funcset_extend_domainI: "\ y \ T; f \ extensional_funcset S T\ \ f(x := y) \ extensional_funcset (insert x S) T" bulwahn@40631: unfolding extensional_funcset_def extensional_def by auto bulwahn@40631: bulwahn@40631: lemma extensional_funcset_restrict_domain: bulwahn@40631: "x \ S \ f \ extensional_funcset (insert x S) T \ f(x := undefined) \ extensional_funcset S T" bulwahn@40631: unfolding extensional_funcset_def extensional_def by auto bulwahn@40631: bulwahn@40631: lemma extensional_funcset_extend_domain_eq: bulwahn@40631: assumes "x \ S" bulwahn@40631: shows bulwahn@40631: "extensional_funcset (insert x S) T = (\(y, g). g(x := y)) ` {(y, g). y \ T \ g \ extensional_funcset S T}" bulwahn@40631: using assms bulwahn@40631: proof - bulwahn@40631: { bulwahn@40631: fix f bulwahn@40631: assume "f : extensional_funcset (insert x S) T" bulwahn@40631: from this assms have "f : (%(y, g). g(x := y)) ` (T <*> extensional_funcset S T)" bulwahn@40631: unfolding image_iff bulwahn@40631: apply (rule_tac x="(f x, f(x := undefined))" in bexI) bulwahn@40631: apply (auto intro: extensional_funcset_extend_domainI extensional_funcset_restrict_domain extensional_funcset_mem) done bulwahn@40631: } bulwahn@40631: moreover bulwahn@40631: { bulwahn@40631: fix f bulwahn@40631: assume "f : (%(y, g). g(x := y)) ` (T <*> extensional_funcset S T)" bulwahn@40631: from this assms have "f : extensional_funcset (insert x S) T" bulwahn@40631: by (auto intro: extensional_funcset_extend_domainI) bulwahn@40631: } bulwahn@40631: ultimately show ?thesis by auto bulwahn@40631: qed bulwahn@40631: bulwahn@40631: lemma extensional_funcset_fun_upd_restricts_rangeI: "\ y \ S. f x \ f y ==> f : extensional_funcset (insert x S) T ==> f(x := undefined) : extensional_funcset S (T - {f x})" bulwahn@40631: unfolding extensional_funcset_def extensional_def bulwahn@40631: apply auto bulwahn@40631: apply (case_tac "x = xa") bulwahn@40631: apply auto done bulwahn@40631: bulwahn@40631: lemma extensional_funcset_fun_upd_extends_rangeI: bulwahn@40631: assumes "a \ T" "f : extensional_funcset S (T - {a})" bulwahn@40631: shows "f(x := a) : extensional_funcset (insert x S) T" bulwahn@40631: using assms unfolding extensional_funcset_def extensional_def by auto bulwahn@40631: bulwahn@40631: subsubsection {* Injective Extensional Function Spaces *} bulwahn@40631: bulwahn@40631: lemma extensional_funcset_fun_upd_inj_onI: bulwahn@40631: assumes "f \ extensional_funcset S (T - {a})" "inj_on f S" bulwahn@40631: shows "inj_on (f(x := a)) S" bulwahn@40631: using assms unfolding extensional_funcset_def by (auto intro!: inj_on_fun_updI) bulwahn@40631: bulwahn@40631: lemma extensional_funcset_extend_domain_inj_on_eq: bulwahn@40631: assumes "x \ S" bulwahn@40631: shows"{f. f \ extensional_funcset (insert x S) T \ inj_on f (insert x S)} = bulwahn@40631: (%(y, g). g(x:=y)) ` {(y, g). y \ T \ g \ extensional_funcset S (T - {y}) \ inj_on g S}" bulwahn@40631: proof - bulwahn@40631: from assms show ?thesis bulwahn@40631: apply auto bulwahn@40631: apply (auto intro: extensional_funcset_fun_upd_inj_onI extensional_funcset_fun_upd_extends_rangeI) bulwahn@40631: apply (auto simp add: image_iff inj_on_def) bulwahn@40631: apply (rule_tac x="xa x" in exI) bulwahn@40631: apply (auto intro: extensional_funcset_mem) bulwahn@40631: apply (rule_tac x="xa(x := undefined)" in exI) bulwahn@40631: apply (auto intro!: extensional_funcset_fun_upd_restricts_rangeI) bulwahn@40631: apply (auto dest!: extensional_funcset_mem split: split_if_asm) bulwahn@40631: done bulwahn@40631: qed bulwahn@40631: bulwahn@40631: lemma extensional_funcset_extend_domain_inj_onI: bulwahn@40631: assumes "x \ S" bulwahn@40631: shows "inj_on (\(y, g). g(x := y)) {(y, g). y \ T \ g \ extensional_funcset S (T - {y}) \ inj_on g S}" bulwahn@40631: proof - bulwahn@40631: from assms show ?thesis bulwahn@40631: apply (auto intro!: inj_onI) bulwahn@40631: apply (metis fun_upd_same) bulwahn@40631: by (metis assms extensional_funcset_arb fun_upd_triv fun_upd_upd) bulwahn@40631: qed bulwahn@40631: bulwahn@40631: bulwahn@40631: subsubsection {* Cardinality *} bulwahn@40631: bulwahn@40631: lemma card_extensional_funcset: bulwahn@40631: assumes "finite S" bulwahn@40631: shows "card (extensional_funcset S T) = (card T) ^ (card S)" bulwahn@40631: using assms bulwahn@40631: proof (induct rule: finite_induct) bulwahn@40631: case empty bulwahn@40631: show ?case bulwahn@40631: by (auto simp add: extensional_funcset_empty_domain) bulwahn@40631: next bulwahn@40631: case (insert x S) bulwahn@40631: { bulwahn@40631: fix g g' y y' bulwahn@40631: assume assms: "g \ extensional_funcset S T" bulwahn@40631: "g' \ extensional_funcset S T" bulwahn@40631: "y \ T" "y' \ T" bulwahn@40631: "g(x := y) = g'(x := y')" bulwahn@40631: from this have "y = y'" bulwahn@40631: by (metis fun_upd_same) bulwahn@40631: have "g = g'" bulwahn@40631: by (metis assms(1) assms(2) assms(5) extensional_funcset_arb fun_upd_triv fun_upd_upd insert(2)) bulwahn@40631: from `y = y'` `g = g'` have "y = y' & g = g'" by simp bulwahn@40631: } bulwahn@40631: from this have "inj_on (\(y, g). g (x := y)) (T \ extensional_funcset S T)" bulwahn@40631: by (auto intro: inj_onI) bulwahn@40631: from this insert.hyps show ?case bulwahn@40631: by (simp add: extensional_funcset_extend_domain_eq card_image card_cartesian_product) bulwahn@40631: qed bulwahn@40631: bulwahn@40631: lemma finite_extensional_funcset: bulwahn@40631: assumes "finite S" "finite T" bulwahn@40631: shows "finite (extensional_funcset S T)" bulwahn@40631: proof - bulwahn@40631: from card_extensional_funcset[OF assms(1), of T] assms(2) bulwahn@40631: have "(card (extensional_funcset S T) \ 0) \ (S \ {} \ T = {})" bulwahn@40631: by auto bulwahn@40631: from this show ?thesis bulwahn@40631: proof bulwahn@40631: assume "card (extensional_funcset S T) \ 0" bulwahn@40631: from this show ?thesis bulwahn@40631: by (auto intro: card_ge_0_finite) bulwahn@40631: next bulwahn@40631: assume "S \ {} \ T = {}" bulwahn@40631: from this show ?thesis bulwahn@40631: by (auto simp add: extensional_funcset_empty_range) bulwahn@40631: qed bulwahn@40631: qed bulwahn@40631: paulson@13586: end