diff --git a/src/HOL/Library/FuncSet.thy b/src/HOL/Library/FuncSet.thy --- a/src/HOL/Library/FuncSet.thy +++ b/src/HOL/Library/FuncSet.thy @@ -1,720 +1,717 @@ (* Title: HOL/Library/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 funcset_to_empty_iff: "A \ {} = (if A={} then UNIV else {})" by auto lemma Pi_eq_empty[simp]: "(\ x \ A. B x) = {} \ (\x\A. B x = {})" proof - have "\x\A. B x = {}" if "\f. \y. y \ A \ f y \ B y" using that [of "\u. SOME y. y \ B u"] some_in_eq by blast then show ?thesis by force qed 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 (metis PiE fun_upd_apply) by force 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" 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) lemma sum_restrict' [simp]: "sum' (\i\I. g i) I = sum' (\i. g i) I" by (simp add: sum.G_def conj_commute cong: conj_cong) lemma prod_restrict' [simp]: "prod' (\i\I. g i) I = prod' (\i. g i) I" by (simp add: prod.G_def conj_commute cong: conj_cong) 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 restrict_PiE_iff [simp]: "restrict f I \ Pi\<^sub>E I X \ (\i \ I. f i \ X i)" + by (simp add: PiE_iff) + lemma ext_funcset_to_sing_iff [simp]: "A \\<^sub>E {a} = {\x\A. a}" by (auto simp: PiE_def Pi_iff extensionalityI) 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 + by (auto split: if_split_asm) 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 lemma subset_PiE: "PiE I S \ PiE I T \ PiE I S = {} \ (\i \ I. S i \ T i)" (is "?lhs \ _ \ ?rhs") proof (cases "PiE I S = {}") case False moreover have "?lhs = ?rhs" proof assume L: ?lhs have "\i. i\I \ S i \ {}" using False PiE_eq_empty_iff by blast with L show ?rhs by (simp add: PiE_Int PiE_eq_iff inf.absorb_iff2) qed auto ultimately show ?thesis by simp qed simp lemma PiE_eq: "PiE I S = PiE I T \ PiE I S = {} \ PiE I T = {} \ (\i \ I. S i = T i)" by (auto simp: PiE_eq_iff PiE_eq_empty_iff) lemma PiE_UNIV [simp]: "PiE UNIV (\i. UNIV) = UNIV" by blast lemma image_projection_PiE: "(\f. f i) ` (PiE I S) = (if PiE I S = {} then {} else if i \ I then S i else {undefined})" proof - have "(\f. f i) ` Pi\<^sub>E I S = S i" if "i \ I" "f \ PiE I S" for f using that apply auto by (rule_tac x="(\k. if k=i then x else f k)" in image_eqI) auto moreover have "(\f. f i) ` Pi\<^sub>E I S = {undefined}" if "f \ PiE I S" "i \ I" for f using that by (blast intro: PiE_arb [OF that, symmetric]) ultimately show ?thesis by auto qed lemma PiE_singleton: assumes "f \ extensional A" shows "PiE A (\x. {f x}) = {f}" proof - { fix g assume "g \ PiE A (\x. {f x})" hence "g x = f x" for x using assms by (cases "x \ A") (auto simp: extensional_def) hence "g = f" by (simp add: fun_eq_iff) } thus ?thesis using assms by (auto simp: extensional_def) qed lemma PiE_eq_singleton: "(\\<^sub>E i\I. S i) = {\i\I. f i} \ (\i\I. S i = {f i})" by (metis (mono_tags, lifting) PiE_eq PiE_singleton insert_not_empty restrict_apply' restrict_extensional) lemma PiE_over_singleton_iff: "(\\<^sub>E x\{a}. B x) = (\b \ B a. {\x \ {a}. b})" apply (auto simp: PiE_iff split: if_split_asm) apply (metis (no_types, lifting) extensionalityI restrict_apply' restrict_extensional singletonD) done lemma all_PiE_elements: "(\z \ PiE I S. \i \ I. P i (z i)) \ PiE I S = {} \ (\i \ I. \x \ S i. P i x)" (is "?lhs = ?rhs") proof (cases "PiE I S = {}") case False then obtain f where f: "\i. i \ I \ f i \ S i" by fastforce show ?thesis proof assume L: ?lhs have "P i x" if "i \ I" "x \ S i" for i x proof - have "(\j \ I. if j=i then x else f j) \ PiE I S" by (simp add: f that(2)) then have "P i ((\j \ I. if j=i then x else f j) i)" using L that(1) by blast with that show ?thesis by simp qed then show ?rhs by (simp add: False) qed fastforce qed simp lemma PiE_ext: "\x \ PiE k s; y \ PiE k s; \i. i \ k \ x i = y i\ \ x = y" by (metis ext PiE_E) 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 \Misc properties of functions, composition and restriction from HOL Light\ lemma function_factors_left_gen: "(\x y. P x \ P y \ g x = g y \ f x = f y) \ (\h. \x. P x \ f x = h(g x))" (is "?lhs = ?rhs") proof assume L: ?lhs then show ?rhs apply (rule_tac x="f \ inv_into (Collect P) g" in exI) unfolding o_def by (metis (mono_tags, hide_lams) f_inv_into_f imageI inv_into_into mem_Collect_eq) qed auto lemma function_factors_left: "(\x y. (g x = g y) \ (f x = f y)) \ (\h. f = h \ g)" using function_factors_left_gen [of "\x. True" g f] unfolding o_def by blast lemma function_factors_right_gen: "(\x. P x \ (\y. g y = f x)) \ (\h. \x. P x \ f x = g(h x))" by metis lemma function_factors_right: "(\x. \y. g y = f x) \ (\h. f = g \ h)" unfolding o_def by metis lemma restrict_compose_right: "restrict (g \ restrict f S) S = restrict (g \ f) S" by auto lemma restrict_compose_left: "f ` S \ T \ restrict (restrict g T \ f) S = restrict (g \ f) S" by fastforce 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 subsection \The pigeonhole principle\ text \ An alternative formulation of this is that for a function mapping a finite set \A\ of cardinality \m\ to a finite set \B\ of cardinality \n\, there exists an element \y \ B\ that is hit at least $\lceil \frac{m}{n}\rceil$ times. However, since we do not have real numbers or rounding yet, we state it in the following equivalent form: \ lemma pigeonhole_card: assumes "f \ A \ B" "finite A" "finite B" "B \ {}" shows "\y\B. card (f -` {y} \ A) * card B \ card A" proof - from assms have "card B > 0" by auto define M where "M = Max ((\y. card (f -` {y} \ A)) ` B)" have "A = (\y\B. f -` {y} \ A)" using assms by auto also have "card \ = (\i\B. card (f -` {i} \ A))" using assms by (subst card_UN_disjoint) auto also have "\ \ (\i\B. M)" unfolding M_def using assms by (intro sum_mono Max.coboundedI) auto also have "\ = card B * M" by simp finally have "M * card B \ card A" by (simp add: mult_ac) moreover have "M \ (\y. card (f -` {y} \ A)) ` B" unfolding M_def using assms \B \ {}\ by (intro Max_in) auto ultimately show ?thesis by blast qed end diff --git a/src/HOL/Metis_Examples/Tarski.thy b/src/HOL/Metis_Examples/Tarski.thy --- a/src/HOL/Metis_Examples/Tarski.thy +++ b/src/HOL/Metis_Examples/Tarski.thy @@ -1,1035 +1,1028 @@ (* Title: HOL/Metis_Examples/Tarski.thy Author: Lawrence C. Paulson, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen Metis example featuring the full theorem of Tarski. *) section \Metis Example Featuring the Full Theorem of Tarski\ theory Tarski imports Main "HOL-Library.FuncSet" begin declare [[metis_new_skolem]] (*Many of these higher-order problems appear to be impossible using the current linkup. They often seem to need either higher-order unification or explicit reasoning about connectives such as conjunction. The numerous set comprehensions are to blame.*) record 'a potype = pset :: "'a set" order :: "('a * 'a) set" definition monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" where "monotone f A r == \x\A. \y\A. (x, y) \ r --> ((f x), (f y)) \ r" definition least :: "['a => bool, 'a potype] => 'a" where "least P po \ SOME x. x \ pset po \ P x \ (\y \ pset po. P y \ (x,y) \ order po)" definition greatest :: "['a => bool, 'a potype] => 'a" where "greatest P po \ SOME x. x \ pset po \ P x \ (\y \ pset po. P y \ (y,x) \ order po)" definition lub :: "['a set, 'a potype] => 'a" where "lub S po == least (\x. \y\S. (y,x) \ order po) po" definition glb :: "['a set, 'a potype] => 'a" where "glb S po \ greatest (\x. \y\S. (x,y) \ order po) po" definition isLub :: "['a set, 'a potype, 'a] => bool" where "isLub S po \ \L. (L \ pset po \ (\y\S. (y,L) \ order po) \ (\z\pset po. (\y\S. (y,z) \ order po) \ (L,z) \ order po))" definition isGlb :: "['a set, 'a potype, 'a] => bool" where "isGlb S po \ \G. (G \ pset po \ (\y\S. (G,y) \ order po) \ (\z \ pset po. (\y\S. (z,y) \ order po) \ (z,G) \ order po))" definition "fix" :: "[('a => 'a), 'a set] => 'a set" where "fix f A \ {x. x \ A \ f x = x}" definition interval :: "[('a*'a) set,'a, 'a ] => 'a set" where "interval r a b == {x. (a,x) \ r & (x,b) \ r}" definition Bot :: "'a potype => 'a" where "Bot po == least (\x. True) po" definition Top :: "'a potype => 'a" where "Top po == greatest (\x. True) po" definition PartialOrder :: "('a potype) set" where "PartialOrder == {P. refl_on (pset P) (order P) & antisym (order P) & trans (order P)}" definition CompleteLattice :: "('a potype) set" where "CompleteLattice == {cl. cl \ PartialOrder \ (\S. S \ pset cl \ (\L. isLub S cl L)) \ (\S. S \ pset cl \ (\G. isGlb S cl G))}" definition induced :: "['a set, ('a * 'a) set] => ('a *'a)set" where "induced A r \ {(a,b). a \ A \ b \ A \ (a,b) \ r}" definition sublattice :: "('a potype * 'a set)set" where "sublattice \ SIGMA cl : CompleteLattice. {S. S \ pset cl \ \pset = S, order = induced S (order cl)\ \ CompleteLattice}" abbreviation sublattice_syntax :: "['a set, 'a potype] => bool" ("_ <<= _" [51, 50] 50) where "S <<= cl \ S \ sublattice `` {cl}" definition dual :: "'a potype => 'a potype" where "dual po == (| pset = pset po, order = converse (order po) |)" locale PO = fixes cl :: "'a potype" and A :: "'a set" and r :: "('a * 'a) set" assumes cl_po: "cl \ PartialOrder" defines A_def: "A == pset cl" and r_def: "r == order cl" locale CL = PO + assumes cl_co: "cl \ CompleteLattice" definition CLF_set :: "('a potype * ('a => 'a)) set" where "CLF_set = (SIGMA cl: CompleteLattice. {f. f \ pset cl \ pset cl \ monotone f (pset cl) (order cl)})" locale CLF = CL + fixes f :: "'a => 'a" and P :: "'a set" assumes f_cl: "(cl,f) \ CLF_set" (*was the equivalent "f : CLF``{cl}"*) defines P_def: "P == fix f A" locale Tarski = CLF + fixes Y :: "'a set" and intY1 :: "'a set" and v :: "'a" assumes Y_ss: "Y \ P" defines intY1_def: "intY1 == interval r (lub Y cl) (Top cl)" and v_def: "v == glb {x. ((\x \ intY1. f x) x, x) \ induced intY1 r \ x \ intY1} \pset=intY1, order=induced intY1 r\" subsection \Partial Order\ lemma (in PO) PO_imp_refl_on: "refl_on A r" apply (insert cl_po) apply (simp add: PartialOrder_def A_def r_def) done lemma (in PO) PO_imp_sym: "antisym r" apply (insert cl_po) apply (simp add: PartialOrder_def r_def) done lemma (in PO) PO_imp_trans: "trans r" apply (insert cl_po) apply (simp add: PartialOrder_def r_def) done lemma (in PO) reflE: "x \ A ==> (x, x) \ r" apply (insert cl_po) apply (simp add: PartialOrder_def refl_on_def A_def r_def) done lemma (in PO) antisymE: "[| (a, b) \ r; (b, a) \ r |] ==> a = b" apply (insert cl_po) apply (simp add: PartialOrder_def antisym_def r_def) done lemma (in PO) transE: "[| (a, b) \ r; (b, c) \ r|] ==> (a,c) \ r" apply (insert cl_po) apply (simp add: PartialOrder_def r_def) apply (unfold trans_def, fast) done lemma (in PO) monotoneE: "[| monotone f A r; x \ A; y \ A; (x, y) \ r |] ==> (f x, f y) \ r" by (simp add: monotone_def) lemma (in PO) po_subset_po: "S \ A ==> (| pset = S, order = induced S r |) \ PartialOrder" apply (simp (no_asm) add: PartialOrder_def) apply auto \ \refl\ apply (simp add: refl_on_def induced_def) apply (blast intro: reflE) \ \antisym\ apply (simp add: antisym_def induced_def) apply (blast intro: antisymE) \ \trans\ apply (simp add: trans_def induced_def) apply (blast intro: transE) done lemma (in PO) indE: "[| (x, y) \ induced S r; S \ A |] ==> (x, y) \ r" by (simp add: add: induced_def) lemma (in PO) indI: "[| (x, y) \ r; x \ S; y \ S |] ==> (x, y) \ induced S r" by (simp add: add: induced_def) lemma (in CL) CL_imp_ex_isLub: "S \ A ==> \L. isLub S cl L" apply (insert cl_co) apply (simp add: CompleteLattice_def A_def) done declare (in CL) cl_co [simp] lemma isLub_lub: "(\L. isLub S cl L) = isLub S cl (lub S cl)" by (simp add: lub_def least_def isLub_def some_eq_ex [symmetric]) lemma isGlb_glb: "(\G. isGlb S cl G) = isGlb S cl (glb S cl)" by (simp add: glb_def greatest_def isGlb_def some_eq_ex [symmetric]) lemma isGlb_dual_isLub: "isGlb S cl = isLub S (dual cl)" by (simp add: isLub_def isGlb_def dual_def converse_unfold) lemma isLub_dual_isGlb: "isLub S cl = isGlb S (dual cl)" by (simp add: isLub_def isGlb_def dual_def converse_unfold) lemma (in PO) dualPO: "dual cl \ PartialOrder" apply (insert cl_po) apply (simp add: PartialOrder_def dual_def) done lemma Rdual: "\S. (S \ A -->( \L. isLub S (| pset = A, order = r|) L)) ==> \S. (S \ A --> (\G. isGlb S (| pset = A, order = r|) G))" apply safe apply (rule_tac x = "lub {y. y \ A & (\k \ S. (y, k) \ r)} (|pset = A, order = r|) " in exI) apply (drule_tac x = "{y. y \ A & (\k \ S. (y,k) \ r) }" in spec) apply (drule mp, fast) apply (simp add: isLub_lub isGlb_def) apply (simp add: isLub_def, blast) done lemma lub_dual_glb: "lub S cl = glb S (dual cl)" by (simp add: lub_def glb_def least_def greatest_def dual_def converse_unfold) lemma glb_dual_lub: "glb S cl = lub S (dual cl)" by (simp add: lub_def glb_def least_def greatest_def dual_def converse_unfold) lemma CL_subset_PO: "CompleteLattice \ PartialOrder" by (simp add: PartialOrder_def CompleteLattice_def, fast) lemmas CL_imp_PO = CL_subset_PO [THEN subsetD] declare PO.PO_imp_refl_on [OF PO.intro [OF CL_imp_PO], simp] declare PO.PO_imp_sym [OF PO.intro [OF CL_imp_PO], simp] declare PO.PO_imp_trans [OF PO.intro [OF CL_imp_PO], simp] lemma (in CL) CO_refl_on: "refl_on A r" by (rule PO_imp_refl_on) lemma (in CL) CO_antisym: "antisym r" by (rule PO_imp_sym) lemma (in CL) CO_trans: "trans r" by (rule PO_imp_trans) lemma CompleteLatticeI: "[| po \ PartialOrder; (\S. S \ pset po --> (\L. isLub S po L)); (\S. S \ pset po --> (\G. isGlb S po G))|] ==> po \ CompleteLattice" apply (unfold CompleteLattice_def, blast) done lemma (in CL) CL_dualCL: "dual cl \ CompleteLattice" apply (insert cl_co) apply (simp add: CompleteLattice_def dual_def) apply (fold dual_def) apply (simp add: isLub_dual_isGlb [symmetric] isGlb_dual_isLub [symmetric] dualPO) done lemma (in PO) dualA_iff: "pset (dual cl) = pset cl" by (simp add: dual_def) lemma (in PO) dualr_iff: "((x, y) \ (order(dual cl))) = ((y, x) \ order cl)" by (simp add: dual_def) lemma (in PO) monotone_dual: "monotone f (pset cl) (order cl) ==> monotone f (pset (dual cl)) (order(dual cl))" by (simp add: monotone_def dualA_iff dualr_iff) lemma (in PO) interval_dual: "[| x \ A; y \ A|] ==> interval r x y = interval (order(dual cl)) y x" apply (simp add: interval_def dualr_iff) apply (fold r_def, fast) done lemma (in PO) interval_not_empty: "[| trans r; interval r a b \ {} |] ==> (a, b) \ r" apply (simp add: interval_def) apply (unfold trans_def, blast) done lemma (in PO) interval_imp_mem: "x \ interval r a b ==> (a, x) \ r" by (simp add: interval_def) lemma (in PO) left_in_interval: "[| a \ A; b \ A; interval r a b \ {} |] ==> a \ interval r a b" apply (simp (no_asm_simp) add: interval_def) apply (simp add: PO_imp_trans interval_not_empty) apply (simp add: reflE) done lemma (in PO) right_in_interval: "[| a \ A; b \ A; interval r a b \ {} |] ==> b \ interval r a b" apply (simp (no_asm_simp) add: interval_def) apply (simp add: PO_imp_trans interval_not_empty) apply (simp add: reflE) done subsection \sublattice\ lemma (in PO) sublattice_imp_CL: "S <<= cl ==> (| pset = S, order = induced S r |) \ CompleteLattice" by (simp add: sublattice_def CompleteLattice_def A_def r_def) lemma (in CL) sublatticeI: "[| S \ A; (| pset = S, order = induced S r |) \ CompleteLattice |] ==> S <<= cl" by (simp add: sublattice_def A_def r_def) subsection \lub\ lemma (in CL) lub_unique: "[| S \ A; isLub S cl x; isLub S cl L|] ==> x = L" apply (rule antisymE) apply (auto simp add: isLub_def r_def) done lemma (in CL) lub_upper: "[|S \ A; x \ S|] ==> (x, lub S cl) \ r" apply (rule CL_imp_ex_isLub [THEN exE], assumption) apply (unfold lub_def least_def) apply (rule some_equality [THEN ssubst]) apply (simp add: isLub_def) apply (simp add: lub_unique A_def isLub_def) apply (simp add: isLub_def r_def) done lemma (in CL) lub_least: "[| S \ A; L \ A; \x \ S. (x,L) \ r |] ==> (lub S cl, L) \ r" apply (rule CL_imp_ex_isLub [THEN exE], assumption) apply (unfold lub_def least_def) apply (rule_tac s=x in some_equality [THEN ssubst]) apply (simp add: isLub_def) apply (simp add: lub_unique A_def isLub_def) apply (simp add: isLub_def r_def A_def) done lemma (in CL) lub_in_lattice: "S \ A ==> lub S cl \ A" apply (rule CL_imp_ex_isLub [THEN exE], assumption) apply (unfold lub_def least_def) apply (subst some_equality) apply (simp add: isLub_def) prefer 2 apply (simp add: isLub_def A_def) apply (simp add: lub_unique A_def isLub_def) done lemma (in CL) lubI: "[| S \ A; L \ A; \x \ S. (x,L) \ r; \z \ A. (\y \ S. (y,z) \ r) --> (L,z) \ r |] ==> L = lub S cl" apply (rule lub_unique, assumption) apply (simp add: isLub_def A_def r_def) apply (unfold isLub_def) apply (rule conjI) apply (fold A_def r_def) apply (rule lub_in_lattice, assumption) apply (simp add: lub_upper lub_least) done lemma (in CL) lubIa: "[| S \ A; isLub S cl L |] ==> L = lub S cl" by (simp add: lubI isLub_def A_def r_def) lemma (in CL) isLub_in_lattice: "isLub S cl L ==> L \ A" by (simp add: isLub_def A_def) lemma (in CL) isLub_upper: "[|isLub S cl L; y \ S|] ==> (y, L) \ r" by (simp add: isLub_def r_def) lemma (in CL) isLub_least: "[| isLub S cl L; z \ A; \y \ S. (y, z) \ r|] ==> (L, z) \ r" by (simp add: isLub_def A_def r_def) lemma (in CL) isLubI: "\L \ A; \y \ S. (y, L) \ r; (\z \ A. (\y \ S. (y, z) \ r) \ (L, z) \ r)\ \ isLub S cl L" by (simp add: isLub_def A_def r_def) subsection \glb\ lemma (in CL) glb_in_lattice: "S \ A ==> glb S cl \ A" apply (subst glb_dual_lub) apply (simp add: A_def) apply (rule dualA_iff [THEN subst]) apply (rule CL.lub_in_lattice) apply (rule CL.intro) apply (rule PO.intro) apply (rule dualPO) apply (rule CL_axioms.intro) apply (rule CL_dualCL) apply (simp add: dualA_iff) done lemma (in CL) glb_lower: "[|S \ A; x \ S|] ==> (glb S cl, x) \ r" apply (subst glb_dual_lub) apply (simp add: r_def) apply (rule dualr_iff [THEN subst]) apply (rule CL.lub_upper) apply (rule CL.intro) apply (rule PO.intro) apply (rule dualPO) apply (rule CL_axioms.intro) apply (rule CL_dualCL) apply (simp add: dualA_iff A_def, assumption) done text \ Reduce the sublattice property by using substructural properties; abandoned see \Tarski_4.ML\. \ declare (in CLF) f_cl [simp] lemma (in CLF) [simp]: "f \ pset cl \ pset cl \ monotone f (pset cl) (order cl)" proof - have "\u v. (v, u) \ CLF_set \ u \ {R \ pset v \ pset v. monotone R (pset v) (order v)}" unfolding CLF_set_def using SigmaE2 by blast hence F1: "\u v. (v, u) \ CLF_set \ u \ pset v \ pset v \ monotone u (pset v) (order v)" using CollectE by blast hence "Tarski.monotone f (pset cl) (order cl)" by (metis f_cl) hence "(cl, f) \ CLF_set \ Tarski.monotone f (pset cl) (order cl)" by (metis f_cl) thus "f \ pset cl \ pset cl \ Tarski.monotone f (pset cl) (order cl)" using F1 by metis qed lemma (in CLF) f_in_funcset: "f \ A \ A" by (simp add: A_def) lemma (in CLF) monotone_f: "monotone f A r" by (simp add: A_def r_def) (*never proved, 2007-01-22*) declare (in CLF) CLF_set_def [simp] CL_dualCL [simp] monotone_dual [simp] dualA_iff [simp] lemma (in CLF) CLF_dual: "(dual cl, f) \ CLF_set" apply (simp del: dualA_iff) apply (simp) done declare (in CLF) CLF_set_def[simp del] CL_dualCL[simp del] monotone_dual[simp del] dualA_iff[simp del] subsection \fixed points\ lemma fix_subset: "fix f A \ A" by (auto simp add: fix_def) lemma fix_imp_eq: "x \ fix f A ==> f x = x" by (simp add: fix_def) lemma fixf_subset: "[| A \ B; x \ fix (\y \ A. f y) A |] ==> x \ fix f B" by (simp add: fix_def, auto) subsection \lemmas for Tarski, lub\ (*never proved, 2007-01-22*) declare CL.lub_least[intro] CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.transE[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] lemma (in CLF) lubH_le_flubH: "H = {x. (x, f x) \ r & x \ A} ==> (lub H cl, f (lub H cl)) \ r" apply (rule lub_least, fast) apply (rule f_in_funcset [THEN funcset_mem]) apply (rule lub_in_lattice, fast) \ \\\x:H. (x, f (lub H r)) \ r\\ apply (rule ballI) (*never proved, 2007-01-22*) apply (rule transE) \ \instantiates \(x, ?z) \ order cl to (x, f x)\,\ \ \because of the definition of \H\\ apply fast \ \so it remains to show \(f x, f (lub H cl)) \ r\\ apply (rule_tac f = "f" in monotoneE) apply (rule monotone_f, fast) apply (rule lub_in_lattice, fast) apply (rule lub_upper, fast) apply assumption done declare CL.lub_least[rule del] CLF.f_in_funcset[rule del] funcset_mem[rule del] CL.lub_in_lattice[rule del] PO.transE[rule del] PO.monotoneE[rule del] CLF.monotone_f[rule del] CL.lub_upper[rule del] (*never proved, 2007-01-22*) declare CLF.f_in_funcset[intro] funcset_mem[intro] CL.lub_in_lattice[intro] PO.monotoneE[intro] CLF.monotone_f[intro] CL.lub_upper[intro] CLF.lubH_le_flubH[simp] lemma (in CLF) flubH_le_lubH: "[| H = {x. (x, f x) \ r & x \ A} |] ==> (f (lub H cl), lub H cl) \ r" apply (rule lub_upper, fast) apply (rule_tac t = "H" in ssubst, assumption) apply (rule CollectI) by (metis (lifting) CO_refl_on lubH_le_flubH monotone_def monotone_f refl_onD1 refl_onD2) declare CLF.f_in_funcset[rule del] funcset_mem[rule del] CL.lub_in_lattice[rule del] PO.monotoneE[rule del] CLF.monotone_f[rule del] CL.lub_upper[rule del] CLF.lubH_le_flubH[simp del] (*never proved, 2007-01-22*) (* Single-step version fails. The conjecture clauses refer to local abstraction functions (Frees). *) lemma (in CLF) lubH_is_fixp: "H = {x. (x, f x) \ r & x \ A} ==> lub H cl \ fix f A" apply (simp add: fix_def) apply (rule conjI) proof - assume A1: "H = {x. (x, f x) \ r \ x \ A}" have F1: "\u v. v \ u \ u" by (metis Int_commute Int_lower1) have "{R. (R, f R) \ r} \ {R. R \ A} = H" using A1 by (metis Collect_conj_eq) hence "H \ {R. R \ A}" using F1 by metis hence "H \ A" by (metis Collect_mem_eq) hence "lub H cl \ A" by (metis lub_in_lattice) thus "lub {x. (x, f x) \ r \ x \ A} cl \ A" using A1 by metis next assume A1: "H = {x. (x, f x) \ r \ x \ A}" have F1: "\v. {R. R \ v} = v" by (metis Collect_mem_eq) have F2: "\w u. {R. R \ u \ R \ w} = u \ w" by (metis Collect_conj_eq Collect_mem_eq) have F3: "\x v. {R. v R \ x} = v -` x" by (metis vimage_def) hence F4: "A \ (\R. (R, f R)) -` r = H" using A1 by auto hence F5: "(f (lub H cl), lub H cl) \ r" by (metis A1 flubH_le_lubH) have F6: "(lub H cl, f (lub H cl)) \ r" by (metis A1 lubH_le_flubH) have "(lub H cl, f (lub H cl)) \ r \ f (lub H cl) = lub H cl" using F5 by (metis antisymE) hence "f (lub H cl) = lub H cl" using F6 by metis thus "H = {x. (x, f x) \ r \ x \ A} \ f (lub {x. (x, f x) \ r \ x \ A} cl) = lub {x. (x, f x) \ r \ x \ A} cl" by metis qed lemma (in CLF) (*lubH_is_fixp:*) "H = {x. (x, f x) \ r & x \ A} ==> lub H cl \ fix f A" apply (simp add: fix_def) apply (rule conjI) apply (metis CO_refl_on lubH_le_flubH refl_onD1) apply (metis antisymE flubH_le_lubH lubH_le_flubH) done lemma (in CLF) fix_in_H: "[| H = {x. (x, f x) \ r & x \ A}; x \ P |] ==> x \ H" by (simp add: P_def fix_imp_eq [of _ f A] reflE CO_refl_on fix_subset [of f A, THEN subsetD]) lemma (in CLF) fixf_le_lubH: "H = {x. (x, f x) \ r & x \ A} ==> \x \ fix f A. (x, lub H cl) \ r" apply (rule ballI) apply (rule lub_upper, fast) apply (rule fix_in_H) apply (simp_all add: P_def) done lemma (in CLF) lubH_least_fixf: "H = {x. (x, f x) \ r & x \ A} ==> \L. (\y \ fix f A. (y,L) \ r) --> (lub H cl, L) \ r" apply (metis P_def lubH_is_fixp) done subsection \Tarski fixpoint theorem 1, first part\ declare CL.lubI[intro] fix_subset[intro] CL.lub_in_lattice[intro] CLF.fixf_le_lubH[simp] CLF.lubH_least_fixf[simp] lemma (in CLF) T_thm_1_lub: "lub P cl = lub {x. (x, f x) \ r & x \ A} cl" (*sledgehammer;*) apply (rule sym) apply (simp add: P_def) apply (rule lubI) apply (simp add: fix_subset) using fix_subset lubH_is_fixp apply fastforce apply (simp add: fixf_le_lubH) using lubH_is_fixp apply blast done declare CL.lubI[rule del] fix_subset[rule del] CL.lub_in_lattice[rule del] CLF.fixf_le_lubH[simp del] CLF.lubH_least_fixf[simp del] (*never proved, 2007-01-22*) declare glb_dual_lub[simp] PO.dualA_iff[intro] CLF.lubH_is_fixp[intro] PO.dualPO[intro] CL.CL_dualCL[intro] PO.dualr_iff[simp] lemma (in CLF) glbH_is_fixp: "H = {x. (f x, x) \ r & x \ A} ==> glb H cl \ P" \ \Tarski for glb\ (*sledgehammer;*) apply (simp add: glb_dual_lub P_def A_def r_def) apply (rule dualA_iff [THEN subst]) apply (rule CLF.lubH_is_fixp) apply (rule CLF.intro) apply (rule CL.intro) apply (rule PO.intro) apply (rule dualPO) apply (rule CL_axioms.intro) apply (rule CL_dualCL) apply (rule CLF_axioms.intro) apply (rule CLF_dual) apply (simp add: dualr_iff dualA_iff) done declare glb_dual_lub[simp del] PO.dualA_iff[rule del] CLF.lubH_is_fixp[rule del] PO.dualPO[rule del] CL.CL_dualCL[rule del] PO.dualr_iff[simp del] (*never proved, 2007-01-22*) lemma (in CLF) T_thm_1_glb: "glb P cl = glb {x. (f x, x) \ r & x \ A} cl" (*sledgehammer;*) apply (simp add: glb_dual_lub P_def A_def r_def) apply (rule dualA_iff [THEN subst]) (*never proved, 2007-01-22*) (*sledgehammer;*) apply (simp add: CLF.T_thm_1_lub [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, OF dualPO CL_dualCL] dualPO CL_dualCL CLF_dual dualr_iff) done subsection \interval\ declare (in CLF) CO_refl_on[simp] refl_on_def [simp] lemma (in CLF) rel_imp_elem: "(x, y) \ r ==> x \ A" by (metis CO_refl_on refl_onD1) declare (in CLF) CO_refl_on[simp del] refl_on_def [simp del] declare (in CLF) rel_imp_elem[intro] declare interval_def [simp] lemma (in CLF) interval_subset: "[| a \ A; b \ A |] ==> interval r a b \ A" by (metis CO_refl_on interval_imp_mem refl_onD refl_onD2 rel_imp_elem subset_eq) declare (in CLF) rel_imp_elem[rule del] declare interval_def [simp del] lemma (in CLF) intervalI: "[| (a, x) \ r; (x, b) \ r |] ==> x \ interval r a b" by (simp add: interval_def) lemma (in CLF) interval_lemma1: "[| S \ interval r a b; x \ S |] ==> (a, x) \ r" by (unfold interval_def, fast) lemma (in CLF) interval_lemma2: "[| S \ interval r a b; x \ S |] ==> (x, b) \ r" by (unfold interval_def, fast) lemma (in CLF) a_less_lub: "[| S \ A; S \ {}; \x \ S. (a,x) \ r; \y \ S. (y, L) \ r |] ==> (a,L) \ r" by (blast intro: transE) lemma (in CLF) glb_less_b: "[| S \ A; S \ {}; \x \ S. (x,b) \ r; \y \ S. (G, y) \ r |] ==> (G,b) \ r" by (blast intro: transE) lemma (in CLF) S_intv_cl: "[| a \ A; b \ A; S \ interval r a b |]==> S \ A" by (simp add: subset_trans [OF _ interval_subset]) lemma (in CLF) L_in_interval: "[| a \ A; b \ A; S \ interval r a b; S \ {}; isLub S cl L; interval r a b \ {} |] ==> L \ interval r a b" (*WON'T TERMINATE apply (metis CO_trans intervalI interval_lemma1 interval_lemma2 isLub_least isLub_upper subset_empty subset_iff trans_def) *) apply (rule intervalI) apply (rule a_less_lub) prefer 2 apply assumption apply (simp add: S_intv_cl) apply (rule ballI) apply (simp add: interval_lemma1) apply (simp add: isLub_upper) \ \\(L, b) \ r\\ apply (simp add: isLub_least interval_lemma2) done (*never proved, 2007-01-22*) lemma (in CLF) G_in_interval: "[| a \ A; b \ A; interval r a b \ {}; S \ interval r a b; isGlb S cl G; S \ {} |] ==> G \ interval r a b" apply (simp add: interval_dual) apply (simp add: CLF.L_in_interval [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualA_iff A_def dualPO CL_dualCL CLF_dual isGlb_dual_isLub) done lemma (in CLF) intervalPO: "[| a \ A; b \ A; interval r a b \ {} |] ==> (| pset = interval r a b, order = induced (interval r a b) r |) \ PartialOrder" proof - assume A1: "a \ A" assume "b \ A" hence "\u. u \ A \ interval r u b \ A" by (metis interval_subset) hence "interval r a b \ A" using A1 by metis hence "interval r a b \ A" by metis thus ?thesis by (metis po_subset_po) qed lemma (in CLF) intv_CL_lub: "[| a \ A; b \ A; interval r a b \ {} |] ==> \S. S \ interval r a b --> (\L. isLub S (| pset = interval r a b, order = induced (interval r a b) r |) L)" apply (intro strip) apply (frule S_intv_cl [THEN CL_imp_ex_isLub]) prefer 2 apply assumption apply assumption apply (erule exE) \ \define the lub for the interval as\ apply (rule_tac x = "if S = {} then a else L" in exI) apply (simp (no_asm_simp) add: isLub_def split del: if_split) apply (intro impI conjI) \ \\(if S = {} then a else L) \ interval r a b\\ apply (simp add: CL_imp_PO L_in_interval) apply (simp add: left_in_interval) \ \lub prop 1\ apply (case_tac "S = {}") \ \\S = {}, y \ S = False => everything\\ apply fast \ \\S \ {}\\ apply simp \ \\\y:S. (y, L) \ induced (interval r a b) r\\ apply (rule ballI) apply (simp add: induced_def L_in_interval) apply (rule conjI) apply (rule subsetD) apply (simp add: S_intv_cl, assumption) apply (simp add: isLub_upper) \ \\\z:interval r a b. (\y:S. (y, z) \ induced (interval r a b) r \ (if S = {} then a else L, z) \ induced (interval r a b) r\\ apply (rule ballI) apply (rule impI) apply (case_tac "S = {}") \ \\S = {}\\ apply simp apply (simp add: induced_def interval_def) apply (rule conjI) apply (rule reflE, assumption) apply (rule interval_not_empty) apply (rule CO_trans) apply (simp add: interval_def) \ \\S \ {}\\ apply simp apply (simp add: induced_def L_in_interval) apply (rule isLub_least, assumption) apply (rule subsetD) prefer 2 apply assumption apply (simp add: S_intv_cl, fast) done lemmas (in CLF) intv_CL_glb = intv_CL_lub [THEN Rdual] (*never proved, 2007-01-22*) lemma (in CLF) interval_is_sublattice: "[| a \ A; b \ A; interval r a b \ {} |] ==> interval r a b <<= cl" (*sledgehammer *) apply (rule sublatticeI) apply (simp add: interval_subset) (*never proved, 2007-01-22*) (*sledgehammer *) apply (rule CompleteLatticeI) apply (simp add: intervalPO) apply (simp add: intv_CL_lub) apply (simp add: intv_CL_glb) done lemmas (in CLF) interv_is_compl_latt = interval_is_sublattice [THEN sublattice_imp_CL] subsection \Top and Bottom\ lemma (in CLF) Top_dual_Bot: "Top cl = Bot (dual cl)" by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff) lemma (in CLF) Bot_dual_Top: "Bot cl = Top (dual cl)" by (simp add: Top_def Bot_def least_def greatest_def dualA_iff dualr_iff) lemma (in CLF) Bot_in_lattice: "Bot cl \ A" (*sledgehammer; *) apply (simp add: Bot_def least_def) apply (rule_tac a="glb A cl" in someI2) apply (simp_all add: glb_in_lattice glb_lower r_def [symmetric] A_def [symmetric]) done (*first proved 2007-01-25 after relaxing relevance*) lemma (in CLF) Top_in_lattice: "Top cl \ A" (*sledgehammer;*) apply (simp add: Top_dual_Bot A_def) (*first proved 2007-01-25 after relaxing relevance*) (*sledgehammer*) apply (rule dualA_iff [THEN subst]) apply (blast intro!: CLF.Bot_in_lattice [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualPO CL_dualCL CLF_dual) done lemma (in CLF) Top_prop: "x \ A ==> (x, Top cl) \ r" apply (simp add: Top_def greatest_def) apply (rule_tac a="lub A cl" in someI2) apply (rule someI2) apply (simp_all add: lub_in_lattice lub_upper r_def [symmetric] A_def [symmetric]) done (*never proved, 2007-01-22*) lemma (in CLF) Bot_prop: "x \ A ==> (Bot cl, x) \ r" (*sledgehammer*) apply (simp add: Bot_dual_Top r_def) apply (rule dualr_iff [THEN subst]) apply (simp add: CLF.Top_prop [of _ f, OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] dualA_iff A_def dualPO CL_dualCL CLF_dual) done lemma (in CLF) Top_intv_not_empty: "x \ A ==> interval r x (Top cl) \ {}" apply (metis Top_in_lattice Top_prop empty_iff intervalI reflE) done lemma (in CLF) Bot_intv_not_empty: "x \ A ==> interval r (Bot cl) x \ {}" apply (metis Bot_prop ex_in_conv intervalI reflE rel_imp_elem) done subsection \fixed points form a partial order\ lemma (in CLF) fixf_po: "(| pset = P, order = induced P r|) \ PartialOrder" by (simp add: P_def fix_subset po_subset_po) (*first proved 2007-01-25 after relaxing relevance*) declare (in Tarski) P_def[simp] Y_ss [simp] declare fix_subset [intro] subset_trans [intro] lemma (in Tarski) Y_subset_A: "Y \ A" (*sledgehammer*) apply (rule subset_trans [OF _ fix_subset]) apply (rule Y_ss [simplified P_def]) done declare (in Tarski) P_def[simp del] Y_ss [simp del] declare fix_subset [rule del] subset_trans [rule del] lemma (in Tarski) lubY_in_A: "lub Y cl \ A" by (rule Y_subset_A [THEN lub_in_lattice]) (*never proved, 2007-01-22*) lemma (in Tarski) lubY_le_flubY: "(lub Y cl, f (lub Y cl)) \ r" (*sledgehammer*) apply (rule lub_least) apply (rule Y_subset_A) apply (rule f_in_funcset [THEN funcset_mem]) apply (rule lubY_in_A) \ \\Y \ P ==> f x = x\\ apply (rule ballI) (*sledgehammer *) apply (rule_tac t = "x" in fix_imp_eq [THEN subst]) apply (erule Y_ss [simplified P_def, THEN subsetD]) \ \\reduce (f x, f (lub Y cl)) \ r to (x, lub Y cl) \ r\ by monotonicity\ (*sledgehammer*) apply (rule_tac f = "f" in monotoneE) apply (rule monotone_f) apply (simp add: Y_subset_A [THEN subsetD]) apply (rule lubY_in_A) apply (simp add: lub_upper Y_subset_A) done (*first proved 2007-01-25 after relaxing relevance*) lemma (in Tarski) intY1_subset: "intY1 \ A" (*sledgehammer*) apply (unfold intY1_def) apply (rule interval_subset) apply (rule lubY_in_A) apply (rule Top_in_lattice) done lemmas (in Tarski) intY1_elem = intY1_subset [THEN subsetD] (*never proved, 2007-01-22*) lemma (in Tarski) intY1_f_closed: "x \ intY1 \ f x \ intY1" (*sledgehammer*) apply (simp add: intY1_def interval_def) apply (rule conjI) apply (rule transE) apply (rule lubY_le_flubY) \ \\(f (lub Y cl), f x) \ r\\ (*sledgehammer [has been proved before now...]*) apply (rule_tac f=f in monotoneE) apply (rule monotone_f) apply (rule lubY_in_A) apply (simp add: intY1_def interval_def intY1_elem) apply (simp add: intY1_def interval_def) \ \\(f x, Top cl) \ r\\ apply (rule Top_prop) apply (rule f_in_funcset [THEN funcset_mem]) apply (simp add: intY1_def interval_def intY1_elem) done lemma (in Tarski) intY1_func: "(\x \ intY1. f x) \ intY1 \ intY1" -apply (rule restrict_in_funcset) -apply (metis intY1_f_closed restrict_in_funcset) +apply (rule restrictI) +apply (metis intY1_f_closed) done lemma (in Tarski) intY1_mono: "monotone (\x \ intY1. f x) intY1 (induced intY1 r)" (*sledgehammer *) apply (auto simp add: monotone_def induced_def intY1_f_closed) apply (blast intro: intY1_elem monotone_f [THEN monotoneE]) done (*proof requires relaxing relevance: 2007-01-25*) lemma (in Tarski) intY1_is_cl: "(| pset = intY1, order = induced intY1 r |) \ CompleteLattice" (*sledgehammer*) apply (unfold intY1_def) apply (rule interv_is_compl_latt) apply (rule lubY_in_A) apply (rule Top_in_lattice) apply (rule Top_intv_not_empty) apply (rule lubY_in_A) done (*never proved, 2007-01-22*) lemma (in Tarski) v_in_P: "v \ P" (*sledgehammer*) apply (unfold P_def) apply (rule_tac A = "intY1" in fixf_subset) apply (rule intY1_subset) apply (simp add: CLF.glbH_is_fixp [OF CLF.intro, OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified] v_def CL_imp_PO intY1_is_cl CLF_set_def intY1_func intY1_mono) done lemma (in Tarski) z_in_interval: "[| z \ P; \y\Y. (y, z) \ induced P r |] ==> z \ intY1" (*sledgehammer *) apply (unfold intY1_def P_def) apply (rule intervalI) prefer 2 apply (erule fix_subset [THEN subsetD, THEN Top_prop]) apply (rule lub_least) apply (rule Y_subset_A) apply (fast elim!: fix_subset [THEN subsetD]) apply (simp add: induced_def) done lemma (in Tarski) f'z_in_int_rel: "[| z \ P; \y\Y. (y, z) \ induced P r |] ==> ((\x \ intY1. f x) z, z) \ induced intY1 r" using P_def fix_imp_eq indI intY1_elem reflE z_in_interval by fastforce (*never proved, 2007-01-22*) lemma (in Tarski) tarski_full_lemma: "\L. isLub Y (| pset = P, order = induced P r |) L" apply (rule_tac x = "v" in exI) apply (simp add: isLub_def) \ \\v \ P\\ apply (simp add: v_in_P) apply (rule conjI) (*sledgehammer*) \ \\v\ is lub\ \ \\1. \y:Y. (y, v) \ induced P r\\ apply (rule ballI) apply (simp add: induced_def subsetD v_in_P) apply (rule conjI) apply (erule Y_ss [THEN subsetD]) apply (rule_tac b = "lub Y cl" in transE) apply (rule lub_upper) apply (rule Y_subset_A, assumption) apply (rule_tac b = "Top cl" in interval_imp_mem) apply (simp add: v_def) apply (fold intY1_def) apply (rule CL.glb_in_lattice [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]) apply (simp add: CL_imp_PO intY1_is_cl, force) \ \\v\ is LEAST ub\ apply clarify apply (rule indI) prefer 3 apply assumption prefer 2 apply (simp add: v_in_P) apply (unfold v_def) (*never proved, 2007-01-22*) (*sledgehammer*) apply (rule indE) apply (rule_tac [2] intY1_subset) (*never proved, 2007-01-22*) (*sledgehammer*) apply (rule CL.glb_lower [OF CL.intro, OF PO.intro CL_axioms.intro, OF _ intY1_is_cl, simplified]) apply (simp add: CL_imp_PO intY1_is_cl) apply force apply (simp add: induced_def intY1_f_closed z_in_interval) apply (simp add: P_def fix_imp_eq [of _ f A] reflE fix_subset [of f A, THEN subsetD]) done lemma CompleteLatticeI_simp: "[| (| pset = A, order = r |) \ PartialOrder; \S. S \ A --> (\L. isLub S (| pset = A, order = r |) L) |] ==> (| pset = A, order = r |) \ CompleteLattice" by (simp add: CompleteLatticeI Rdual) (*never proved, 2007-01-22*) declare (in CLF) fixf_po[intro] P_def [simp] A_def [simp] r_def [simp] Tarski.tarski_full_lemma [intro] cl_po [intro] cl_co [intro] CompleteLatticeI_simp [intro] theorem (in CLF) Tarski_full: "(| pset = P, order = induced P r|) \ CompleteLattice" + using A_def CLF_axioms P_def Tarski.intro Tarski_axioms.intro fixf_po r_def by blast (*sledgehammer*) -apply (rule CompleteLatticeI_simp) -apply (rule fixf_po, clarify) -(*never proved, 2007-01-22*) -(*sledgehammer*) -apply (simp add: P_def A_def r_def) -apply (blast intro!: Tarski.tarski_full_lemma [OF Tarski.intro, OF CLF.intro Tarski_axioms.intro, - OF CL.intro CLF_axioms.intro, OF PO.intro CL_axioms.intro] cl_po cl_co f_cl) -done declare (in CLF) fixf_po [rule del] P_def [simp del] A_def [simp del] r_def [simp del] Tarski.tarski_full_lemma [rule del] cl_po [rule del] cl_co [rule del] CompleteLatticeI_simp [rule del] end