# HG changeset patch # User haftmann # Date 1312868657 -7200 # Node ID c0847967a25a77f914148f7905aa6ab3cd111235 # Parent 730f7cced3a6ab5b0270f29f6f2419c3c3fdcd8c# Parent a65e26f1427bffff0a58b852603a78a144276aed merged diff -r 730f7cced3a6 -r c0847967a25a NEWS --- a/NEWS Mon Aug 08 19:26:53 2011 -0700 +++ b/NEWS Tue Aug 09 07:44:17 2011 +0200 @@ -68,6 +68,8 @@ boolean algebras and complete linear orders. Lemmas Inf_less_iff, less_Sup_iff, INF_less_iff, less_SUP_iff now reside in class complete_linorder. Changes proposition of lemmas Inf_fun_def, Sup_fun_def, Inf_apply, Sup_apply. +Redundant lemmas Inf_singleton, Sup_singleton, Inf_binary and Sup_binary have +been discarded. More consistent and less misunderstandable names: INFI_def ~> INF_def SUPR_def ~> SUP_def diff -r 730f7cced3a6 -r c0847967a25a src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Mon Aug 08 19:26:53 2011 -0700 +++ b/src/HOL/Complete_Lattice.thy Tue Aug 09 07:44:17 2011 +0200 @@ -126,16 +126,16 @@ lemma le_SUP_I2: "i \ A \ u \ f i \ u \ (\i\A. f i)" using le_SUP_I [of i A f] by auto -lemma le_Inf_iff: "b \ \A \ (\a\A. b \ a)" +lemma le_Inf_iff (*[simp]*): "b \ \A \ (\a\A. b \ a)" by (auto intro: Inf_greatest dest: Inf_lower) -lemma le_INF_iff: "u \ (\i\A. f i) \ (\i \ A. u \ f i)" +lemma le_INF_iff (*[simp]*): "u \ (\i\A. f i) \ (\i\A. u \ f i)" by (auto simp add: INF_def le_Inf_iff) -lemma Sup_le_iff: "\A \ b \ (\a\A. a \ b)" +lemma Sup_le_iff (*[simp]*): "\A \ b \ (\a\A. a \ b)" by (auto intro: Sup_least dest: Sup_upper) -lemma SUP_le_iff: "(\i\A. f i) \ u \ (\i \ A. f i \ u)" +lemma SUP_le_iff (*[simp]*): "(\i\A. f i) \ u \ (\i\A. f i \ u)" by (auto simp add: SUP_def Sup_le_iff) lemma Inf_empty [simp]: @@ -160,22 +160,22 @@ "\UNIV = \" by (auto intro!: antisym Sup_upper) -lemma Inf_insert: "\insert a A = a \ \A" +lemma Inf_insert (*[simp]*): "\insert a A = a \ \A" by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) lemma INF_insert: "(\x\insert a A. f x) = f a \ INFI A f" by (simp add: INF_def Inf_insert) -lemma Sup_insert: "\insert a A = a \ \A" +lemma Sup_insert (*[simp]*): "\insert a A = a \ \A" by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) lemma SUP_insert: "(\x\insert a A. f x) = f a \ SUPR A f" by (simp add: SUP_def Sup_insert) -lemma INF_image: "(\x\f`A. g x) = (\x\A. g (f x))" +lemma INF_image (*[simp]*): "(\x\f`A. g x) = (\x\A. g (f x))" by (simp add: INF_def image_image) -lemma SUP_image: "(\x\f`A. g x) = (\x\A. g (f x))" +lemma SUP_image (*[simp]*): "(\x\f`A. g x) = (\x\A. g (f x))" by (simp add: SUP_def image_image) lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" @@ -184,22 +184,6 @@ lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) -lemma Inf_singleton [simp]: - "\{a} = a" - by (auto intro: antisym Inf_lower Inf_greatest) - -lemma Sup_singleton [simp]: - "\{a} = a" - by (auto intro: antisym Sup_upper Sup_least) - -lemma Inf_binary: - "\{a, b} = a \ b" - by (simp add: Inf_insert) - -lemma Sup_binary: - "\{a, b} = a \ b" - by (simp add: Sup_insert) - lemma Inf_superset_mono: "B \ A \ \A \ \B" by (auto intro: Inf_greatest Inf_lower) @@ -298,7 +282,7 @@ by (rule antisym) (auto intro: le_supI1 le_supI2 le_SUP_I SUP_mono, rule SUP_leI, auto intro: le_supI1 le_supI2 le_SUP_I SUP_mono) -lemma Inf_top_conv [no_atp]: +lemma Inf_top_conv (*[simp]*) [no_atp]: "\A = \ \ (\x\A. x = \)" "\ = \A \ (\x\A. x = \)" proof - @@ -306,7 +290,7 @@ proof assume "\x\A. x = \" then have "A = {} \ A = {\}" by auto - then show "\A = \" by auto + then show "\A = \" by (auto simp add: Inf_insert) next assume "\A = \" show "\x\A. x = \" @@ -320,12 +304,12 @@ then show "\ = \A \ (\x\A. x = \)" by auto qed -lemma INF_top_conv: +lemma INF_top_conv (*[simp]*): "(\x\A. B x) = \ \ (\x\A. B x = \)" "\ = (\x\A. B x) \ (\x\A. B x = \)" by (auto simp add: INF_def Inf_top_conv) -lemma Sup_bot_conv [no_atp]: +lemma Sup_bot_conv (*[simp]*) [no_atp]: "\A = \ \ (\x\A. x = \)" (is ?P) "\ = \A \ (\x\A. x = \)" (is ?Q) proof - @@ -334,7 +318,7 @@ from dual.Inf_top_conv show ?P and ?Q by simp_all qed -lemma SUP_bot_conv: +lemma SUP_bot_conv (*[simp]*): "(\x\A. B x) = \ \ (\x\A. B x = \)" "\ = (\x\A. B x) \ (\x\A. B x = \)" by (auto simp add: SUP_def Sup_bot_conv) @@ -345,10 +329,10 @@ lemma SUP_const [simp]: "A \ {} \ (\i\A. f) = f" by (auto intro: antisym SUP_leI le_SUP_I) -lemma INF_top: "(\x\A. \) = \" +lemma INF_top (*[simp]*): "(\x\A. \) = \" by (cases "A = {}") (simp_all add: INF_empty) -lemma SUP_bot: "(\x\A. \) = \" +lemma SUP_bot (*[simp]*): "(\x\A. \) = \" by (cases "A = {}") (simp_all add: SUP_empty) lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" @@ -381,14 +365,6 @@ "(\y\A. c) = (if A = {} then \ else c)" by (simp add: SUP_empty) -lemma INF_eq: - "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by (simp add: INF_def image_def) - -lemma SUP_eq: - "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by (simp add: SUP_def image_def) - lemma less_INF_D: assumes "y < (\i\A. f i)" "i \ A" shows "y < f i" proof - @@ -407,14 +383,6 @@ finally show "f i < y" . qed -lemma INF_UNIV_range: - "(\x. f x) = \range f" - by (fact INF_def) - -lemma SUP_UNIV_range: - "(\x. f x) = \range f" - by (fact SUP_def) - lemma INF_UNIV_bool_expand: "(\b. A b) = A True \ A False" by (simp add: UNIV_bool INF_empty INF_insert inf_commute) @@ -450,7 +418,7 @@ and prove @{text inf_Sup} and @{text sup_Inf} from that? *} fix a b c from sup_Inf have "a \ \{b, c} = (\d\{b, c}. a \ d)" . - then show "a \ b \ c = (a \ b) \ (a \ c)" by (simp add: INF_def Inf_binary) + then show "a \ b \ c = (a \ b) \ (a \ c)" by (simp add: INF_def Inf_insert) qed lemma Inf_sup: @@ -525,23 +493,23 @@ "class.complete_linorder Sup Inf (op \) (op >) sup inf \ \" by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder) -lemma Inf_less_iff: +lemma Inf_less_iff (*[simp]*): "\S \ a \ (\x\S. x \ a)" unfolding not_le [symmetric] le_Inf_iff by auto -lemma INF_less_iff: +lemma INF_less_iff (*[simp]*): "(\i\A. f i) \ a \ (\x\A. f x \ a)" unfolding INF_def Inf_less_iff by auto -lemma less_Sup_iff: +lemma less_Sup_iff (*[simp]*): "a \ \S \ (\x\S. a \ x)" unfolding not_le [symmetric] Sup_le_iff by auto -lemma less_SUP_iff: +lemma less_SUP_iff (*[simp]*): "a \ (\i\A. f i) \ (\x\A. a \ f x)" unfolding SUP_def less_Sup_iff by auto -lemma Sup_eq_top_iff: +lemma Sup_eq_top_iff (*[simp]*): "\A = \ \ (\x<\. \i\A. x < i)" proof assume *: "\A = \" @@ -563,11 +531,11 @@ qed qed -lemma SUP_eq_top_iff: +lemma SUP_eq_top_iff (*[simp]*): "(\i\A. f i) = \ \ (\x<\. \i\A. x < f i)" unfolding SUP_def Sup_eq_top_iff by auto -lemma Inf_eq_bot_iff: +lemma Inf_eq_bot_iff (*[simp]*): "\A = \ \ (\x>\. \i\A. i < x)" proof - interpret dual: complete_linorder Sup Inf "op \" "op >" sup inf \ \ @@ -575,7 +543,7 @@ from dual.Sup_eq_top_iff show ?thesis . qed -lemma INF_eq_bot_iff: +lemma INF_eq_bot_iff (*[simp]*): "(\i\A. f i) = \ \ (\x>\. \i\A. f i < x)" unfolding INF_def Inf_eq_bot_iff by auto @@ -703,9 +671,6 @@ lemma Inter_greatest: "(\X. X \ A \ C \ X) \ C \ Inter A" by (fact Inf_greatest) -lemma Int_eq_Inter: "A \ B = \{A, B}" - by (fact Inf_binary [symmetric]) - lemma Inter_empty: "\{} = UNIV" by (fact Inf_empty) (* already simp *) @@ -762,34 +727,26 @@ [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}] *} -- {* to avoid eta-contraction of body *} -lemma INTER_eq_Inter_image: - "(\x\A. B x) = \(B`A)" - by (fact INF_def) - -lemma Inter_def: - "\S = (\x\S. x)" - by (simp add: INTER_eq_Inter_image image_def) - -lemma INTER_def: +lemma INTER_eq: "(\x\A. B x) = {y. \x\A. y \ B x}" - by (auto simp add: INTER_eq_Inter_image Inter_eq) + by (auto simp add: INF_def) lemma Inter_image_eq [simp]: "\(B`A) = (\x\A. B x)" by (rule sym) (fact INF_def) lemma INT_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" - by (unfold INTER_def) blast + by (auto simp add: INF_def image_def) lemma INT_I [intro!]: "(\x. x \ A \ b \ B x) \ b \ (\x\A. B x)" - by (unfold INTER_def) blast + by (auto simp add: INF_def image_def) lemma INT_D [elim, Pure.elim]: "b \ (\x\A. B x) \ a \ A \ b \ B a" by auto lemma INT_E [elim]: "b \ (\x\A. B x) \ (b \ B a \ R) \ (a \ A \ R) \ R" -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a\A"}. *} - by (unfold INTER_def) blast + by (auto simp add: INF_def image_def) lemma INT_cong [cong]: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" @@ -808,7 +765,7 @@ by (fact le_INF_I) lemma INT_empty: "(\x\{}. B x) = UNIV" - by (fact INF_empty) (* already simp *) + by (fact INF_empty) lemma INT_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" by (fact INF_absorb) @@ -829,10 +786,6 @@ lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" by (fact INF_constant) -lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" - -- {* Look: it has an \emph{existential} quantifier *} - by (fact INF_eq) - lemma INTER_UNIV_conv [simp]: "(UNIV = (\x\A. B x)) = (\x\A. B x = UNIV)" "((\x\A. B x) = UNIV) = (\x\A. B x = UNIV)" @@ -891,9 +844,6 @@ lemma Union_least: "(\X. X \ A \ X \ C) \ \A \ C" by (fact Sup_least) -lemma Un_eq_Union: "A \ B = \{A, B}" - by blast - lemma Union_empty [simp]: "\{} = {}" by (fact Sup_empty) @@ -966,24 +916,16 @@ [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}] *} -- {* to avoid eta-contraction of body *} -lemma UNION_eq_Union_image: - "(\x\A. B x) = \(B ` A)" - by (fact SUP_def) - -lemma Union_def: - "\S = (\x\S. x)" - by (simp add: UNION_eq_Union_image image_def) - -lemma UNION_def [no_atp]: +lemma UNION_eq [no_atp]: "(\x\A. B x) = {y. \x\A. y \ B x}" - by (auto simp add: UNION_eq_Union_image Union_eq) + by (auto simp add: SUP_def) lemma Union_image_eq [simp]: "\(B ` A) = (\x\A. B x)" - by (rule sym) (fact UNION_eq_Union_image) + by (auto simp add: UNION_eq) lemma UN_iff [simp]: "(b \ (\x\A. B x)) = (\x\A. b \ B x)" - by (unfold UNION_def) blast + by (auto simp add: SUP_def image_def) lemma UN_I [intro]: "a \ A \ b \ B a \ b \ (\x\A. B x)" -- {* The order of the premises presupposes that @{term A} is rigid; @@ -991,7 +933,7 @@ by auto lemma UN_E [elim!]: "b \ (\x\A. B x) \ (\x. x\A \ b \ B x \ R) \ R" - by (unfold UNION_def) blast + by (auto simp add: SUP_def image_def) lemma UN_cong [cong]: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" @@ -1017,21 +959,18 @@ by blast lemma UN_empty [no_atp]: "(\x\{}. B x) = {}" - by (fact SUP_empty) (* already simp *) + by (fact SUP_empty) lemma UN_empty2 [simp]: "(\x\A. {}) = {}" by (fact SUP_bot) -lemma UN_singleton [simp]: "(\x\A. {x}) = A" - by blast - lemma UN_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" by (fact SUP_absorb) lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" by (fact SUP_insert) -lemma UN_Un[simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" +lemma UN_Un [simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" by (fact SUP_union) lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" @@ -1043,9 +982,6 @@ lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" by (fact SUP_constant) -lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by (fact SUP_eq) - lemma image_Union: "f ` \S = (\x\S. f ` x)" by blast @@ -1212,6 +1148,22 @@ text {* Legacy names *} +lemma (in complete_lattice) Inf_singleton [simp]: + "\{a} = a" + by (simp add: Inf_insert) + +lemma (in complete_lattice) Sup_singleton [simp]: + "\{a} = a" + by (simp add: Sup_insert) + +lemma (in complete_lattice) Inf_binary: + "\{a, b} = a \ b" + by (simp add: Inf_insert) + +lemma (in complete_lattice) Sup_binary: + "\{a, b} = a \ b" + by (simp add: Sup_insert) + lemmas (in complete_lattice) INFI_def = INF_def lemmas (in complete_lattice) SUPR_def = SUP_def lemmas (in complete_lattice) le_SUPI = le_SUP_I @@ -1219,12 +1171,65 @@ lemmas (in complete_lattice) le_INFI = le_INF_I lemmas (in complete_lattice) less_INFD = less_INF_D +lemmas INFI_apply = INF_apply +lemmas SUPR_apply = SUP_apply + +text {* Grep and put to news from here *} + +lemma (in complete_lattice) INF_eq: + "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (simp add: INF_def image_def) + +lemma (in complete_lattice) SUP_eq: + "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (simp add: SUP_def image_def) + lemma (in complete_lattice) INF_subset: "B \ A \ INFI A f \ INFI B f" by (rule INF_superset_mono) auto -lemmas INFI_apply = INF_apply -lemmas SUPR_apply = SUP_apply +lemma (in complete_lattice) INF_UNIV_range: + "(\x. f x) = \range f" + by (fact INF_def) + +lemma (in complete_lattice) SUP_UNIV_range: + "(\x. f x) = \range f" + by (fact SUP_def) + +lemma Int_eq_Inter: "A \ B = \{A, B}" + by (simp add: Inf_insert) + +lemma INTER_eq_Inter_image: + "(\x\A. B x) = \(B`A)" + by (fact INF_def) + +lemma Inter_def: + "\S = (\x\S. x)" + by (simp add: INTER_eq_Inter_image image_def) + +lemmas INTER_def = INTER_eq +lemmas UNION_def = UNION_eq + +lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (fact INF_eq) + +lemma Un_eq_Union: "A \ B = \{A, B}" + by blast + +lemma UNION_eq_Union_image: + "(\x\A. B x) = \(B ` A)" + by (fact SUP_def) + +lemma Union_def: + "\S = (\x\S. x)" + by (simp add: UNION_eq_Union_image image_def) + +lemma UN_singleton [simp]: "(\x\A. {x}) = A" + by blast + +lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (fact SUP_eq) + text {* Finally *} diff -r 730f7cced3a6 -r c0847967a25a src/HOL/Import/HOLLight/hollight.imp --- a/src/HOL/Import/HOLLight/hollight.imp Mon Aug 08 19:26:53 2011 -0700 +++ b/src/HOL/Import/HOLLight/hollight.imp Tue Aug 09 07:44:17 2011 +0200 @@ -591,7 +591,6 @@ "UNIONS_IMAGE" > "HOLLight.hollight.UNIONS_IMAGE" "UNIONS_GSPEC" > "HOLLight.hollight.UNIONS_GSPEC" "UNIONS_2" > "Complete_Lattice.Un_eq_Union" - "UNIONS_1" > "Complete_Lattice.complete_lattice_class.Sup_singleton" "UNIONS_0" > "Complete_Lattice.Union_empty" "UNCURRY_def" > "HOLLight.hollight.UNCURRY_def" "TYDEF_recspace" > "HOLLight.hollight.TYDEF_recspace" @@ -1598,7 +1597,6 @@ "INTERS_IMAGE" > "HOLLight.hollight.INTERS_IMAGE" "INTERS_GSPEC" > "HOLLight.hollight.INTERS_GSPEC" "INTERS_2" > "Complete_Lattice.Int_eq_Inter" - "INTERS_1" > "Complete_Lattice.complete_lattice_class.Inf_singleton" "INTERS_0" > "Complete_Lattice.Inter_empty" "INSERT_UNIV" > "HOLLight.hollight.INSERT_UNIV" "INSERT_UNION_EQ" > "Set.Un_insert_left" diff -r 730f7cced3a6 -r c0847967a25a src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Mon Aug 08 19:26:53 2011 -0700 +++ b/src/HOL/Lattices.thy Tue Aug 09 07:44:17 2011 +0200 @@ -177,10 +177,10 @@ lemma inf_left_commute: "x \ (y \ z) = y \ (x \ z)" by (fact inf.left_commute) -lemma inf_idem: "x \ x = x" +lemma inf_idem (*[simp]*): "x \ x = x" by (fact inf.idem) -lemma inf_left_idem: "x \ (x \ y) = x \ y" +lemma inf_left_idem (*[simp]*): "x \ (x \ y) = x \ y" by (fact inf.left_idem) lemma inf_absorb1: "x \ y \ x \ y = x" @@ -216,10 +216,10 @@ lemma sup_left_commute: "x \ (y \ z) = y \ (x \ z)" by (fact sup.left_commute) -lemma sup_idem: "x \ x = x" +lemma sup_idem (*[simp]*): "x \ x = x" by (fact sup.idem) -lemma sup_left_idem: "x \ (x \ y) = x \ y" +lemma sup_left_idem (*[simp]*): "x \ (x \ y) = x \ y" by (fact sup.left_idem) lemma sup_absorb1: "y \ x \ x \ y = x" @@ -240,10 +240,10 @@ by (rule class.lattice.intro, rule dual_semilattice, rule class.semilattice_sup.intro, rule dual_order) (unfold_locales, auto) -lemma inf_sup_absorb: "x \ (x \ y) = x" +lemma inf_sup_absorb (*[simp]*): "x \ (x \ y) = x" by (blast intro: antisym inf_le1 inf_greatest sup_ge1) -lemma sup_inf_absorb: "x \ (x \ y) = x" +lemma sup_inf_absorb (*[simp]*): "x \ (x \ y) = x" by (blast intro: antisym sup_ge1 sup_least inf_le1) lemmas inf_sup_aci = inf_aci sup_aci @@ -436,11 +436,11 @@ by (rule class.boolean_algebra.intro, rule dual_bounded_lattice, rule dual_distrib_lattice) (unfold_locales, auto simp add: inf_compl_bot sup_compl_top diff_eq) -lemma compl_inf_bot: +lemma compl_inf_bot (*[simp]*): "- x \ x = \" by (simp add: inf_commute inf_compl_bot) -lemma compl_sup_top: +lemma compl_sup_top (*[simp]*): "- x \ x = \" by (simp add: sup_commute sup_compl_top) @@ -522,7 +522,7 @@ then show "- y \ - x" by (simp only: le_iff_inf) qed -lemma compl_le_compl_iff: (* TODO: declare [simp] ? *) +lemma compl_le_compl_iff (*[simp]*): "- x \ - y \ y \ x" by (auto dest: compl_mono)