# HG changeset patch # User haftmann # Date 1291823277 -3600 # Node ID a434f89a9962347a31ca0f38454f99fd608de5fc # Parent 051251fde456575dc8f7f1b6647f85aea26a4cfc# Parent c987429a1298cbb313ece7cb941eb151feb611da merged diff -r 051251fde456 -r a434f89a9962 NEWS --- a/NEWS Wed Dec 08 14:25:08 2010 +0100 +++ b/NEWS Wed Dec 08 16:47:57 2010 +0100 @@ -111,6 +111,17 @@ *** HOL *** +* More canonical naming convention for some fundamental definitions: + + bot_bool_eq ~> bot_bool_def + top_bool_eq ~> top_bool_def + inf_bool_eq ~> inf_bool_def + sup_bool_eq ~> sup_bool_def + bot_fun_eq ~> bot_fun_def + top_fun_eq ~> top_fun_def + inf_fun_eq ~> inf_fun_def + sup_fun_eq ~> sup_fun_def + * Quickcheck now by default uses exhaustive testing instead of random testing. Random testing can be invoked by quickcheck[random], exhaustive testing by quickcheck[exhaustive]. diff -r 051251fde456 -r a434f89a9962 src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Wed Dec 08 14:25:08 2010 +0100 +++ b/src/HOL/Complete_Lattice.thy Wed Dec 08 16:47:57 2010 +0100 @@ -44,14 +44,22 @@ lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) -lemma Inf_empty: +lemma Inf_empty [simp]: "\{} = \" by (auto intro: antisym Inf_greatest) -lemma Sup_empty: +lemma Sup_empty [simp]: "\{} = \" by (auto intro: antisym Sup_least) +lemma Inf_UNIV [simp]: + "\UNIV = \" + by (simp add: Sup_Inf Sup_empty [symmetric]) + +lemma Sup_UNIV [simp]: + "\UNIV = \" + by (simp add: Inf_Sup Inf_empty [symmetric]) + lemma Inf_insert: "\insert a A = a \ \A" by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) @@ -74,19 +82,21 @@ "\{a, b} = a \ b" by (simp add: Sup_empty Sup_insert) -lemma Inf_UNIV: - "\UNIV = bot" - by (simp add: Sup_Inf Sup_empty [symmetric]) - -lemma Sup_UNIV: - "\UNIV = top" - by (simp add: Inf_Sup Inf_empty [symmetric]) +lemma le_Inf_iff: "b \ Inf A \ (\a\A. b \ a)" + by (auto intro: Inf_greatest dest: Inf_lower) lemma Sup_le_iff: "Sup A \ b \ (\a\A. a \ b)" by (auto intro: Sup_least dest: Sup_upper) -lemma le_Inf_iff: "b \ Inf A \ (\a\A. b \ a)" - by (auto intro: Inf_greatest dest: Inf_lower) +lemma Inf_mono: + assumes "\b. b \ B \ \a\A. a \ b" + shows "Inf A \ Inf B" +proof (rule Inf_greatest) + fix b assume "b \ B" + with assms obtain a where "a \ A" and "a \ b" by blast + from `a \ A` have "Inf A \ a" by (rule Inf_lower) + with `a \ b` show "Inf A \ b" by auto +qed lemma Sup_mono: assumes "\a. a \ A \ \b\B. a \ b" @@ -98,43 +108,39 @@ with `a \ b` show "a \ Sup B" by auto qed -lemma Inf_mono: - assumes "\b. b \ B \ \a\A. a \ b" - shows "Inf A \ Inf B" -proof (rule Inf_greatest) - fix b assume "b \ B" - with assms obtain a where "a \ A" and "a \ b" by blast - from `a \ A` have "Inf A \ a" by (rule Inf_lower) - with `a \ b` show "Inf A \ b" by auto -qed +definition INFI :: "'b set \ ('b \ 'a) \ 'a" where + "INFI A f = \ (f ` A)" definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where "SUPR A f = \ (f ` A)" -definition INFI :: "'b set \ ('b \ 'a) \ 'a" where - "INFI A f = \ (f ` A)" - end syntax - "_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10) - "_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 0, 10] 10) - "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) - "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 0, 10] 10) + "_INF1" :: "pttrns \ 'b \ 'b" ("(3INF _./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3INF _:_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3SUP _./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3SUP _:_./ _)" [0, 0, 10] 10) + +syntax (xsymbols) + "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) translations + "INF x y. B" == "INF x. INF y. B" + "INF x. B" == "CONST INFI CONST UNIV (%x. B)" + "INF x. B" == "INF x:CONST UNIV. B" + "INF x:A. B" == "CONST INFI A (%x. B)" "SUP x y. B" == "SUP x. SUP y. B" "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" "SUP x. B" == "SUP x:CONST UNIV. B" "SUP x:A. B" == "CONST SUPR A (%x. B)" - "INF x y. B" == "INF x. INF y. B" - "INF x. B" == "CONST INFI CONST UNIV (%x. B)" - "INF x. B" == "INF x:CONST UNIV. B" - "INF x:A. B" == "CONST INFI A (%x. B)" print_translation {* - [Syntax.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}, - Syntax.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"}] + [Syntax.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"}, + Syntax.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}] *} -- {* to avoid eta-contraction of body *} context complete_lattice @@ -158,79 +164,71 @@ lemma le_INF_iff: "u \ (INF i:A. M i) \ (\i \ A. u \ M i)" unfolding INFI_def by (auto simp add: le_Inf_iff) -lemma SUP_const[simp]: "A \ {} \ (SUP i:A. M) = M" - by (auto intro: antisym SUP_leI le_SUPI) - lemma INF_const[simp]: "A \ {} \ (INF i:A. M) = M" by (auto intro: antisym INF_leI le_INFI) -lemma SUP_mono: - "(\n. n \ A \ \m\B. f n \ g m) \ (SUP n:A. f n) \ (SUP n:B. g n)" - by (force intro!: Sup_mono simp: SUPR_def) +lemma SUP_const[simp]: "A \ {} \ (SUP i:A. M) = M" + by (auto intro: antisym SUP_leI le_SUPI) lemma INF_mono: "(\m. m \ B \ \n\A. f n \ g m) \ (INF n:A. f n) \ (INF n:B. g n)" by (force intro!: Inf_mono simp: INFI_def) -lemma SUP_subset: "A \ B \ SUPR A f \ SUPR B f" - by (intro SUP_mono) auto +lemma SUP_mono: + "(\n. n \ A \ \m\B. f n \ g m) \ (SUP n:A. f n) \ (SUP n:B. g n)" + by (force intro!: Sup_mono simp: SUPR_def) lemma INF_subset: "A \ B \ INFI B f \ INFI A f" by (intro INF_mono) auto -lemma SUP_commute: "(SUP i:A. SUP j:B. f i j) = (SUP j:B. SUP i:A. f i j)" - by (iprover intro: SUP_leI le_SUPI order_trans antisym) +lemma SUP_subset: "A \ B \ SUPR A f \ SUPR B f" + by (intro SUP_mono) auto lemma INF_commute: "(INF i:A. INF j:B. f i j) = (INF j:B. INF i:A. f i j)" by (iprover intro: INF_leI le_INFI order_trans antisym) +lemma SUP_commute: "(SUP i:A. SUP j:B. f i j) = (SUP j:B. SUP i:A. f i j)" + by (iprover intro: SUP_leI le_SUPI order_trans antisym) + end +lemma Inf_less_iff: + fixes a :: "'a\{complete_lattice,linorder}" + shows "Inf S < a \ (\x\S. x < a)" + unfolding not_le[symmetric] le_Inf_iff by auto + lemma less_Sup_iff: fixes a :: "'a\{complete_lattice,linorder}" shows "a < Sup S \ (\x\S. a < x)" unfolding not_le[symmetric] Sup_le_iff by auto -lemma Inf_less_iff: - fixes a :: "'a\{complete_lattice,linorder}" - shows "Inf S < a \ (\x\S. x < a)" - unfolding not_le[symmetric] le_Inf_iff by auto +lemma INF_less_iff: + fixes a :: "'a::{complete_lattice,linorder}" + shows "(INF i:A. f i) < a \ (\x\A. f x < a)" + unfolding INFI_def Inf_less_iff by auto lemma less_SUP_iff: fixes a :: "'a::{complete_lattice,linorder}" shows "a < (SUP i:A. f i) \ (\x\A. a < f x)" unfolding SUPR_def less_Sup_iff by auto -lemma INF_less_iff: - fixes a :: "'a::{complete_lattice,linorder}" - shows "(INF i:A. f i) < a \ (\x\A. f x < a)" - unfolding INFI_def Inf_less_iff by auto - subsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} instantiation bool :: complete_lattice begin definition - Inf_bool_def: "\A \ (\x\A. x)" + "\A \ (\x\A. x)" definition - Sup_bool_def: "\A \ (\x\A. x)" + "\A \ (\x\A. x)" instance proof qed (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) end -lemma Inf_empty_bool [simp]: - "\{}" - unfolding Inf_bool_def by auto - -lemma not_Sup_empty_bool [simp]: - "\ \{}" - unfolding Sup_bool_def by auto - -lemma INFI_bool_eq: +lemma INFI_bool_eq [simp]: "INFI = Ball" proof (rule ext)+ fix A :: "'a set" @@ -239,7 +237,7 @@ by (auto simp add: Ball_def INFI_def Inf_bool_def) qed -lemma SUPR_bool_eq: +lemma SUPR_bool_eq [simp]: "SUPR = Bex" proof (rule ext)+ fix A :: "'a set" @@ -252,36 +250,226 @@ begin definition - Inf_fun_def: "\A = (\x. \{y. \f\A. y = f x})" + "\A = (\x. \{y. \f\A. y = f x})" + +lemma Inf_apply: + "(\A) x = \{y. \f\A. y = f x}" + by (simp add: Inf_fun_def) definition - Sup_fun_def: "\A = (\x. \{y. \f\A. y = f x})" + "\A = (\x. \{y. \f\A. y = f x})" + +lemma Sup_apply: + "(\A) x = \{y. \f\A. y = f x}" + by (simp add: Sup_fun_def) instance proof -qed (auto simp add: Inf_fun_def Sup_fun_def le_fun_def +qed (auto simp add: le_fun_def Inf_apply Sup_apply intro: Inf_lower Sup_upper Inf_greatest Sup_least) end -lemma SUPR_fun_expand: - fixes f :: "'a \ 'b \ 'c\{complete_lattice}" - shows "(SUP y:A. f y) = (\x. SUP y:A. f y x)" - by (auto intro!: arg_cong[where f=Sup] ext[where 'a='b] - simp: SUPR_def Sup_fun_def) +lemma INFI_apply: + "(\y\A. f y) x = (\y\A. f y x)" + by (auto intro: arg_cong [of _ _ Inf] simp add: INFI_def Inf_apply) + +lemma SUPR_apply: + "(\y\A. f y) x = (\y\A. f y x)" + by (auto intro: arg_cong [of _ _ Sup] simp add: SUPR_def Sup_apply) + + +subsection {* Inter *} + +abbreviation Inter :: "'a set set \ 'a set" where + "Inter S \ \S" + +notation (xsymbols) + Inter ("\_" [90] 90) + +lemma Inter_eq: + "\A = {x. \B \ A. x \ B}" +proof (rule set_eqI) + fix x + have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" + by auto + then show "x \ \A \ x \ {x. \B \ A. x \ B}" + by (simp add: Inf_fun_def Inf_bool_def) (simp add: mem_def) +qed + +lemma Inter_iff [simp,no_atp]: "(A : Inter C) = (ALL X:C. A:X)" + by (unfold Inter_eq) blast + +lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C" + by (simp add: Inter_eq) + +text {* + \medskip A ``destruct'' rule -- every @{term X} in @{term C} + contains @{term A} as an element, but @{prop "A:X"} can hold when + @{prop "X:C"} does not! This rule is analogous to @{text spec}. +*} + +lemma InterD [elim, Pure.elim]: "A : Inter C ==> X:C ==> A:X" + by auto + +lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R" + -- {* ``Classical'' elimination rule -- does not require proving + @{prop "X:C"}. *} + by (unfold Inter_eq) blast + +lemma Inter_lower: "B \ A ==> Inter A \ B" + by blast + +lemma Inter_subset: + "[| !!X. X \ A ==> X \ B; A ~= {} |] ==> \A \ B" + by blast + +lemma Inter_greatest: "(!!X. X \ A ==> C \ X) ==> C \ Inter A" + by (iprover intro: InterI subsetI dest: subsetD) + +lemma Int_eq_Inter: "A \ B = \{A, B}" + by blast + +lemma Inter_empty [simp]: "\{} = UNIV" + by (fact Inf_empty) + +lemma Inter_UNIV [simp]: "\UNIV = {}" + by blast + +lemma Inter_insert [simp]: "\(insert a B) = a \ \B" + by blast + +lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" + by blast + +lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" + by blast + +lemma Inter_UNIV_conv [simp,no_atp]: + "(\A = UNIV) = (\x\A. x = UNIV)" + "(UNIV = \A) = (\x\A. x = UNIV)" + by blast+ + +lemma Inter_anti_mono: "B \ A ==> \A \ \B" + by blast + + +subsection {* Intersections of families *} + +abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where + "INTER \ INFI" + +syntax + "_INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) + "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 0, 10] 10) + +syntax (xsymbols) + "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) + "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) + +syntax (latex output) + "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) + "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) -lemma INFI_fun_expand: - fixes f :: "'a \ 'b \ 'c\{complete_lattice}" - shows "(INF y:A. f y) x = (INF y:A. f y x)" - by (auto intro!: arg_cong[where f=Inf] ext[where 'a='b] - simp: INFI_def Inf_fun_def) +translations + "INT x y. B" == "INT x. INT y. B" + "INT x. B" == "CONST INTER CONST UNIV (%x. B)" + "INT x. B" == "INT x:CONST UNIV. B" + "INT x:A. B" == "CONST INTER A (%x. B)" + +print_translation {* + [Syntax.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 INFI_def) + +lemma Inter_def: + "\S = (\x\S. x)" + by (simp add: INTER_eq_Inter_image image_def) + +lemma INTER_def: + "(\x\A. B x) = {y. \x\A. y \ B x}" + by (auto simp add: INTER_eq_Inter_image Inter_eq) + +lemma Inter_image_eq [simp]: + "\(B`A) = (\x\A. B x)" + by (rule sym) (fact INTER_eq_Inter_image) + +lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)" + by (unfold INTER_def) blast + +lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)" + by (unfold INTER_def) blast + +lemma INT_D [elim, Pure.elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a" + by auto + +lemma INT_E [elim]: "b : (INT 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 + +lemma INT_cong [cong]: + "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)" + by (simp add: INTER_def) + +lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" + by blast + +lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" + by blast -lemma Inf_empty_fun: - "\{} = (\_. \{})" - by (simp add: Inf_fun_def) +lemma INT_lower: "a \ A ==> (\x\A. B x) \ B a" + by (fact INF_leI) + +lemma INT_greatest: "(!!x. x \ A ==> C \ B x) ==> C \ (\x\A. B x)" + by (fact le_INFI) + +lemma INT_empty [simp]: "(\x\{}. B x) = UNIV" + by blast + +lemma INT_absorb: "k \ I ==> A k \ (\i\I. A i) = (\i\I. A i)" + by blast + +lemma INT_subset_iff: "(B \ (\i\I. A i)) = (\i\I. B \ A i)" + by (fact le_INF_iff) + +lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" + by blast + +lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" + by blast + +lemma INT_insert_distrib: + "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" + by blast -lemma Sup_empty_fun: - "\{} = (\_. \{})" - by (simp add: Sup_fun_def) +lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" + by auto + +lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + -- {* Look: it has an \emph{existential} quantifier *} + by blast + +lemma INTER_UNIV_conv[simp]: + "(UNIV = (INT x:A. B x)) = (\x\A. B x = UNIV)" + "((INT x:A. B x) = UNIV) = (\x\A. B x = UNIV)" +by blast+ + +lemma INT_bool_eq: "(\b::bool. A b) = (A True \ A False)" + by (auto intro: bool_induct) + +lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" + by blast + +lemma INT_anti_mono: + "B \ A ==> (!!x. x \ A ==> f x \ g x) ==> + (\x\A. f x) \ (\x\A. g x)" + -- {* The last inclusion is POSITIVE! *} + by (blast dest: subsetD) + +lemma vimage_INT: "f-`(INT x:A. B x) = (INT x:A. f -` B x)" + by blast subsection {* Union *} @@ -520,200 +708,6 @@ by blast -subsection {* Inter *} - -abbreviation Inter :: "'a set set \ 'a set" where - "Inter S \ \S" - -notation (xsymbols) - Inter ("\_" [90] 90) - -lemma Inter_eq: - "\A = {x. \B \ A. x \ B}" -proof (rule set_eqI) - fix x - have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" - by auto - then show "x \ \A \ x \ {x. \B \ A. x \ B}" - by (simp add: Inf_fun_def Inf_bool_def) (simp add: mem_def) -qed - -lemma Inter_iff [simp,no_atp]: "(A : Inter C) = (ALL X:C. A:X)" - by (unfold Inter_eq) blast - -lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C" - by (simp add: Inter_eq) - -text {* - \medskip A ``destruct'' rule -- every @{term X} in @{term C} - contains @{term A} as an element, but @{prop "A:X"} can hold when - @{prop "X:C"} does not! This rule is analogous to @{text spec}. -*} - -lemma InterD [elim, Pure.elim]: "A : Inter C ==> X:C ==> A:X" - by auto - -lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R" - -- {* ``Classical'' elimination rule -- does not require proving - @{prop "X:C"}. *} - by (unfold Inter_eq) blast - -lemma Inter_lower: "B \ A ==> Inter A \ B" - by blast - -lemma Inter_subset: - "[| !!X. X \ A ==> X \ B; A ~= {} |] ==> \A \ B" - by blast - -lemma Inter_greatest: "(!!X. X \ A ==> C \ X) ==> C \ Inter A" - by (iprover intro: InterI subsetI dest: subsetD) - -lemma Int_eq_Inter: "A \ B = \{A, B}" - by blast - -lemma Inter_empty [simp]: "\{} = UNIV" - by blast - -lemma Inter_UNIV [simp]: "\UNIV = {}" - by blast - -lemma Inter_insert [simp]: "\(insert a B) = a \ \B" - by blast - -lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" - by blast - -lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" - by blast - -lemma Inter_UNIV_conv [simp,no_atp]: - "(\A = UNIV) = (\x\A. x = UNIV)" - "(UNIV = \A) = (\x\A. x = UNIV)" - by blast+ - -lemma Inter_anti_mono: "B \ A ==> \A \ \B" - by blast - - -subsection {* Intersections of families *} - -abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where - "INTER \ INFI" - -syntax - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 0, 10] 10) - -syntax (xsymbols) - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) - -syntax (latex output) - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) - -translations - "INT x y. B" == "INT x. INT y. B" - "INT x. B" == "CONST INTER CONST UNIV (%x. B)" - "INT x. B" == "INT x:CONST UNIV. B" - "INT x:A. B" == "CONST INTER A (%x. B)" - -print_translation {* - [Syntax.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 INFI_def) - -lemma Inter_def: - "\S = (\x\S. x)" - by (simp add: INTER_eq_Inter_image image_def) - -lemma INTER_def: - "(\x\A. B x) = {y. \x\A. y \ B x}" - by (auto simp add: INTER_eq_Inter_image Inter_eq) - -lemma Inter_image_eq [simp]: - "\(B`A) = (\x\A. B x)" - by (rule sym) (fact INTER_eq_Inter_image) - -lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)" - by (unfold INTER_def) blast - -lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)" - by (unfold INTER_def) blast - -lemma INT_D [elim, Pure.elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a" - by auto - -lemma INT_E [elim]: "b : (INT 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 - -lemma INT_cong [cong]: - "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)" - by (simp add: INTER_def) - -lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" - by blast - -lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" - by blast - -lemma INT_lower: "a \ A ==> (\x\A. B x) \ B a" - by (fact INF_leI) - -lemma INT_greatest: "(!!x. x \ A ==> C \ B x) ==> C \ (\x\A. B x)" - by (fact le_INFI) - -lemma INT_empty [simp]: "(\x\{}. B x) = UNIV" - by blast - -lemma INT_absorb: "k \ I ==> A k \ (\i\I. A i) = (\i\I. A i)" - by blast - -lemma INT_subset_iff: "(B \ (\i\I. A i)) = (\i\I. B \ A i)" - by (fact le_INF_iff) - -lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" - by blast - -lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" - by blast - -lemma INT_insert_distrib: - "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" - by blast - -lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" - by auto - -lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" - -- {* Look: it has an \emph{existential} quantifier *} - by blast - -lemma INTER_UNIV_conv[simp]: - "(UNIV = (INT x:A. B x)) = (\x\A. B x = UNIV)" - "((INT x:A. B x) = UNIV) = (\x\A. B x = UNIV)" -by blast+ - -lemma INT_bool_eq: "(\b::bool. A b) = (A True \ A False)" - by (auto intro: bool_induct) - -lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" - by blast - -lemma INT_anti_mono: - "B \ A ==> (!!x. x \ A ==> f x \ g x) ==> - (\x\A. f x) \ (\x\A. g x)" - -- {* The last inclusion is POSITIVE! *} - by (blast dest: subsetD) - -lemma vimage_INT: "f-`(INT x:A. B x) = (INT x:A. f -` B x)" - by blast - - subsection {* Distributive laws *} lemma Int_Union: "A \ \B = (\C\B. A \ C)" @@ -864,12 +858,18 @@ no_notation less_eq (infix "\" 50) and less (infix "\" 50) and + bot ("\") and + top ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) and - top ("\") and - bot ("\") + Sup ("\_" [900] 900) + +no_syntax (xsymbols) + "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) lemmas mem_simps = insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff diff -r 051251fde456 -r a434f89a9962 src/HOL/Inductive.thy --- a/src/HOL/Inductive.thy Wed Dec 08 14:25:08 2010 +0100 +++ b/src/HOL/Inductive.thy Wed Dec 08 16:47:57 2010 +0100 @@ -217,7 +217,8 @@ lemma coinduct3: "[| mono(f); a:X; X \ f(lfp(%x. f(x) Un X Un gfp(f))) |] ==> a : gfp(f)" apply (rule coinduct3_lemma [THEN [2] weak_coinduct]) -apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst], auto) +apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst]) +apply (simp_all) done diff -r 051251fde456 -r a434f89a9962 src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Wed Dec 08 14:25:08 2010 +0100 +++ b/src/HOL/Lattices.thy Wed Dec 08 16:47:57 2010 +0100 @@ -48,8 +48,9 @@ notation less_eq (infix "\" 50) and less (infix "\" 50) and - top ("\") and - bot ("\") + bot ("\") and + top ("\") + class semilattice_inf = order + fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) @@ -587,34 +588,33 @@ begin definition - bool_Compl_def: "uminus = Not" + bool_Compl_def [simp]: "uminus = Not" definition - bool_diff_def: "A - B \ A \ \ B" + bool_diff_def [simp]: "A - B \ A \ \ B" definition - inf_bool_def: "P \ Q \ P \ Q" + [simp]: "P \ Q \ P \ Q" definition - sup_bool_def: "P \ Q \ P \ Q" + [simp]: "P \ Q \ P \ Q" instance proof -qed (simp_all add: inf_bool_def sup_bool_def le_bool_def - bot_bool_def top_bool_def bool_Compl_def bool_diff_def, auto) +qed auto end lemma sup_boolI1: "P \ P \ Q" - by (simp add: sup_bool_def) + by simp lemma sup_boolI2: "Q \ P \ Q" - by (simp add: sup_bool_def) + by simp lemma sup_boolE: "P \ Q \ (P \ R) \ (Q \ R) \ R" - by (auto simp add: sup_bool_def) + by auto subsection {* Fun as lattice *} @@ -623,19 +623,26 @@ begin definition - inf_fun_def: "f \ g = (\x. f x \ g x)" + "f \ g = (\x. f x \ g x)" + +lemma inf_apply: + "(f \ g) x = f x \ g x" + by (simp add: inf_fun_def) definition - sup_fun_def: "f \ g = (\x. f x \ g x)" + "f \ g = (\x. f x \ g x)" + +lemma sup_apply: + "(f \ g) x = f x \ g x" + by (simp add: sup_fun_def) instance proof -qed (simp_all add: le_fun_def inf_fun_def sup_fun_def) +qed (simp_all add: le_fun_def inf_apply sup_apply) end -instance "fun" :: (type, distrib_lattice) distrib_lattice -proof -qed (simp_all add: inf_fun_def sup_fun_def sup_inf_distrib1) +instance "fun" :: (type, distrib_lattice) distrib_lattice proof +qed (rule ext, simp add: sup_inf_distrib1 inf_apply sup_apply) instance "fun" :: (type, bounded_lattice) bounded_lattice .. @@ -645,6 +652,10 @@ definition fun_Compl_def: "- A = (\x. - A x)" +lemma uminus_apply: + "(- A) x = - (A x)" + by (simp add: fun_Compl_def) + instance .. end @@ -655,15 +666,16 @@ definition fun_diff_def: "A - B = (\x. A x - B x)" +lemma minus_apply: + "(A - B) x = A x - B x" + by (simp add: fun_diff_def) + instance .. end -instance "fun" :: (type, boolean_algebra) boolean_algebra -proof -qed (simp_all add: inf_fun_def sup_fun_def bot_fun_def top_fun_def fun_Compl_def fun_diff_def - inf_compl_bot sup_compl_top diff_eq) - +instance "fun" :: (type, boolean_algebra) boolean_algebra proof +qed (rule ext, simp_all add: inf_apply sup_apply bot_apply top_apply uminus_apply minus_apply inf_compl_bot sup_compl_top diff_eq)+ no_notation less_eq (infix "\" 50) and diff -r 051251fde456 -r a434f89a9962 src/HOL/Library/Lattice_Syntax.thy --- a/src/HOL/Library/Lattice_Syntax.thy Wed Dec 08 14:25:08 2010 +0100 +++ b/src/HOL/Library/Lattice_Syntax.thy Wed Dec 08 16:47:57 2010 +0100 @@ -7,11 +7,17 @@ begin notation + bot ("\") and top ("\") and - bot ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\_" [900] 900) and Sup ("\_" [900] 900) +syntax (xsymbols) + "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + end diff -r 051251fde456 -r a434f89a9962 src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Wed Dec 08 14:25:08 2010 +0100 +++ b/src/HOL/Orderings.thy Wed Dec 08 16:47:57 2010 +0100 @@ -1082,14 +1082,14 @@ subsection {* Top and bottom elements *} +class bot = preorder + + fixes bot :: 'a + assumes bot_least [simp]: "bot \ x" + class top = preorder + fixes top :: 'a assumes top_greatest [simp]: "x \ top" -class bot = preorder + - fixes bot :: 'a - assumes bot_least [simp]: "bot \ x" - subsection {* Dense orders *} @@ -1204,50 +1204,50 @@ subsection {* Order on bool *} -instantiation bool :: "{order, top, bot}" +instantiation bool :: "{order, bot, top}" begin definition - le_bool_def: "P \ Q \ P \ Q" + le_bool_def [simp]: "P \ Q \ P \ Q" definition - less_bool_def: "(P\bool) < Q \ \ P \ Q" + [simp]: "(P\bool) < Q \ \ P \ Q" definition - top_bool_def: "top = True" + [simp]: "bot \ False" definition - bot_bool_def: "bot = False" + [simp]: "top \ True" instance proof -qed (auto simp add: bot_bool_def top_bool_def less_bool_def, auto simp add: le_bool_def) +qed auto end lemma le_boolI: "(P \ Q) \ P \ Q" - by (simp add: le_bool_def) + by simp lemma le_boolI': "P \ Q \ P \ Q" - by (simp add: le_bool_def) + by simp lemma le_boolE: "P \ Q \ P \ (Q \ R) \ R" - by (simp add: le_bool_def) + by simp lemma le_boolD: "P \ Q \ P \ Q" - by (simp add: le_bool_def) + by simp lemma bot_boolE: "bot \ P" - by (simp add: bot_bool_def) + by simp lemma top_boolI: top - by (simp add: top_bool_def) + by simp lemma [code]: "False \ b \ True" "True \ b \ b" "False < b \ b" "True < b \ False" - unfolding le_bool_def less_bool_def by simp_all + by simp_all subsection {* Order on functions *} @@ -1259,7 +1259,7 @@ le_fun_def: "f \ g \ (\x. f x \ g x)" definition - less_fun_def: "(f\'a \ 'b) < g \ f \ g \ \ (g \ f)" + "(f\'a \ 'b) < g \ f \ g \ \ (g \ f)" instance .. @@ -1272,26 +1272,34 @@ instance "fun" :: (type, order) order proof qed (auto simp add: le_fun_def intro: antisym ext) +instantiation "fun" :: (type, bot) bot +begin + +definition + "bot = (\x. bot)" + +lemma bot_apply: + "bot x = bot" + by (simp add: bot_fun_def) + +instance proof +qed (simp add: le_fun_def bot_apply) + +end + instantiation "fun" :: (type, top) top begin definition - top_fun_def [no_atp]: "top = (\x. top)" + [no_atp]: "top = (\x. top)" declare top_fun_def_raw [no_atp] +lemma top_apply: + "top x = top" + by (simp add: top_fun_def) + instance proof -qed (simp add: top_fun_def le_fun_def) - -end - -instantiation "fun" :: (type, bot) bot -begin - -definition - bot_fun_def: "bot = (\x. bot)" - -instance proof -qed (simp add: bot_fun_def le_fun_def) +qed (simp add: le_fun_def top_apply) end diff -r 051251fde456 -r a434f89a9962 src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Wed Dec 08 14:25:08 2010 +0100 +++ b/src/HOL/Predicate.thy Wed Dec 08 16:47:57 2010 +0100 @@ -9,12 +9,18 @@ begin notation + bot ("\") and + top ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) and - top ("\") and - bot ("\") + Sup ("\_" [900] 900) + +syntax (xsymbols) + "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) subsection {* Predicates as (complete) lattices *} @@ -86,12 +92,6 @@ subsubsection {* Top and bottom elements *} -lemma top1I [intro!]: "top x" - by (simp add: top_fun_def top_bool_def) - -lemma top2I [intro!]: "top x y" - by (simp add: top_fun_def top_bool_def) - lemma bot1E [no_atp, elim!]: "bot x \ P" by (simp add: bot_fun_def bot_bool_def) @@ -104,6 +104,45 @@ lemma bot_empty_eq2: "bot = (\x y. (x, y) \ {})" by (auto simp add: fun_eq_iff) +lemma top1I [intro!]: "top x" + by (simp add: top_fun_def top_bool_def) + +lemma top2I [intro!]: "top x y" + by (simp add: top_fun_def top_bool_def) + + +subsubsection {* Binary intersection *} + +lemma inf1I [intro!]: "A x ==> B x ==> inf A B x" + by (simp add: inf_fun_def inf_bool_def) + +lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y" + by (simp add: inf_fun_def inf_bool_def) + +lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P" + by (simp add: inf_fun_def inf_bool_def) + +lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P" + by (simp add: inf_fun_def inf_bool_def) + +lemma inf1D1: "inf A B x ==> A x" + by (simp add: inf_fun_def inf_bool_def) + +lemma inf2D1: "inf A B x y ==> A x y" + by (simp add: inf_fun_def inf_bool_def) + +lemma inf1D2: "inf A B x ==> B x" + by (simp add: inf_fun_def inf_bool_def) + +lemma inf2D2: "inf A B x y ==> B x y" + by (simp add: inf_fun_def inf_bool_def) + +lemma inf_Int_eq: "inf (\x. x \ R) (\x. x \ S) = (\x. x \ R \ S)" + by (simp add: inf_fun_def inf_bool_def mem_def) + +lemma inf_Int_eq2 [pred_set_conv]: "inf (\x y. (x, y) \ R) (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" + by (simp add: inf_fun_def inf_bool_def mem_def) + subsubsection {* Binary union *} @@ -143,97 +182,64 @@ by (simp add: sup_fun_def sup_bool_def mem_def) -subsubsection {* Binary intersection *} +subsubsection {* Intersections of families *} -lemma inf1I [intro!]: "A x ==> B x ==> inf A B x" - by (simp add: inf_fun_def inf_bool_def) +lemma INF1_iff: "(INF x:A. B x) b = (ALL x:A. B x b)" + by (simp add: INFI_apply) -lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y" - by (simp add: inf_fun_def inf_bool_def) +lemma INF2_iff: "(INF x:A. B x) b c = (ALL x:A. B x b c)" + by (simp add: INFI_apply) -lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P" - by (simp add: inf_fun_def inf_bool_def) +lemma INF1_I [intro!]: "(!!x. x : A ==> B x b) ==> (INF x:A. B x) b" + by (auto simp add: INFI_apply) -lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P" - by (simp add: inf_fun_def inf_bool_def) +lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c" + by (auto simp add: INFI_apply) -lemma inf1D1: "inf A B x ==> A x" - by (simp add: inf_fun_def inf_bool_def) +lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b" + by (auto simp add: INFI_apply) -lemma inf2D1: "inf A B x y ==> A x y" - by (simp add: inf_fun_def inf_bool_def) +lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c" + by (auto simp add: INFI_apply) -lemma inf1D2: "inf A B x ==> B x" - by (simp add: inf_fun_def inf_bool_def) +lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R" + by (auto simp add: INFI_apply) -lemma inf2D2: "inf A B x y ==> B x y" - by (simp add: inf_fun_def inf_bool_def) +lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R" + by (auto simp add: INFI_apply) -lemma inf_Int_eq: "inf (\x. x \ R) (\x. x \ S) = (\x. x \ R \ S)" - by (simp add: inf_fun_def inf_bool_def mem_def) +lemma INF_INT_eq: "(INF i. (\x. x \ r i)) = (\x. x \ (INT i. r i))" + by (simp add: INFI_apply fun_eq_iff) -lemma inf_Int_eq2 [pred_set_conv]: "inf (\x y. (x, y) \ R) (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" - by (simp add: inf_fun_def inf_bool_def mem_def) +lemma INF_INT_eq2: "(INF i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (INT i. r i))" + by (simp add: INFI_apply fun_eq_iff) subsubsection {* Unions of families *} lemma SUP1_iff: "(SUP x:A. B x) b = (EX x:A. B x b)" - by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast + by (simp add: SUPR_apply) lemma SUP2_iff: "(SUP x:A. B x) b c = (EX x:A. B x b c)" - by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast + by (simp add: SUPR_apply) lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b" - by (auto simp add: SUP1_iff) + by (auto simp add: SUPR_apply) lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c" - by (auto simp add: SUP2_iff) + by (auto simp add: SUPR_apply) lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R" - by (auto simp add: SUP1_iff) + by (auto simp add: SUPR_apply) lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R" - by (auto simp add: SUP2_iff) + by (auto simp add: SUPR_apply) lemma SUP_UN_eq: "(SUP i. (\x. x \ r i)) = (\x. x \ (UN i. r i))" - by (simp add: SUP1_iff fun_eq_iff) + by (simp add: SUPR_apply fun_eq_iff) lemma SUP_UN_eq2: "(SUP i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (UN i. r i))" - by (simp add: SUP2_iff fun_eq_iff) - - -subsubsection {* Intersections of families *} - -lemma INF1_iff: "(INF x:A. B x) b = (ALL x:A. B x b)" - by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast - -lemma INF2_iff: "(INF x:A. B x) b c = (ALL x:A. B x b c)" - by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast - -lemma INF1_I [intro!]: "(!!x. x : A ==> B x b) ==> (INF x:A. B x) b" - by (auto simp add: INF1_iff) - -lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c" - by (auto simp add: INF2_iff) - -lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b" - by (auto simp add: INF1_iff) - -lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c" - by (auto simp add: INF2_iff) - -lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R" - by (auto simp add: INF1_iff) - -lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R" - by (auto simp add: INF2_iff) - -lemma INF_INT_eq: "(INF i. (\x. x \ r i)) = (\x. x \ (INT i. r i))" - by (simp add: INF1_iff fun_eq_iff) - -lemma INF_INT_eq2: "(INF i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (INT i. r i))" - by (simp add: INF2_iff fun_eq_iff) + by (simp add: SUPR_apply fun_eq_iff) subsection {* Predicates as relations *} @@ -493,8 +499,7 @@ by (simp add: minus_pred_def) instance proof -qed (auto intro!: pred_eqI simp add: less_eq_pred_def less_pred_def - fun_Compl_def fun_diff_def bool_Compl_def bool_diff_def) +qed (auto intro!: pred_eqI simp add: less_eq_pred_def less_pred_def uminus_apply minus_apply) end @@ -514,10 +519,10 @@ by (simp add: single_def) definition bind :: "'a pred \ ('a \ 'b pred) \ 'b pred" (infixl "\=" 70) where - "P \= f = (SUPR {x. Predicate.eval P x} f)" + "P \= f = (SUPR {x. eval P x} f)" lemma eval_bind [simp]: - "eval (P \= f) = Predicate.eval (SUPR {x. Predicate.eval P x} f)" + "eval (P \= f) = eval (SUPR {x. eval P x} f)" by (simp add: bind_def) lemma bind_bind: @@ -822,7 +827,7 @@ "single x = Seq (\u. Insert x \)" unfolding Seq_def by simp -primrec "apply" :: "('a \ 'b Predicate.pred) \ 'a seq \ 'b seq" where +primrec "apply" :: "('a \ 'b pred) \ 'a seq \ 'b seq" where "apply f Empty = Empty" | "apply f (Insert x P) = Join (f x) (Join (P \= f) Empty)" | "apply f (Join P xq) = Join (P \= f) (apply f xq)" @@ -972,7 +977,7 @@ "the A = (THE x. eval A x)" lemma the_eqI: - "(THE x. Predicate.eval P x) = x \ Predicate.the P = x" + "(THE x. eval P x) = x \ the P = x" by (simp add: the_def) lemma the_eq [code]: "the A = singleton (\x. not_unique A) A" @@ -1022,14 +1027,20 @@ *} no_notation + bot ("\") and + top ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\_" [900] 900) and Sup ("\_" [900] 900) and - top ("\") and - bot ("\") and bind (infixl "\=" 70) +no_syntax (xsymbols) + "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + hide_type (open) pred seq hide_const (open) Pred eval single bind is_empty singleton if_pred not_pred holds Empty Insert Join Seq member pred_of_seq "apply" adjunct null the_only eq map not_unique the diff -r 051251fde456 -r a434f89a9962 src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Wed Dec 08 14:25:08 2010 +0100 +++ b/src/HOL/Probability/Borel_Space.thy Wed Dec 08 16:47:57 2010 +0100 @@ -1391,7 +1391,7 @@ proof safe fix a have "{x\space M. a < ?sup x} = (\i\A. {x\space M. a < f i x})" - by (auto simp: less_Sup_iff SUPR_def[where 'a=pextreal] SUPR_fun_expand[where 'c=pextreal]) + by (auto simp: less_SUP_iff SUPR_apply) then show "{x\space M. a < ?sup x} \ sets M" using assms by auto qed @@ -1404,7 +1404,7 @@ proof safe fix a have "{x\space M. ?inf x < a} = (\i\A. {x\space M. f i x < a})" - by (auto simp: Inf_less_iff INFI_def[where 'a=pextreal] INFI_fun_expand) + by (auto simp: INF_less_iff INFI_apply) then show "{x\space M. ?inf x < a} \ sets M" using assms by auto qed @@ -1423,11 +1423,20 @@ using assms by auto qed +lemma INFI_fun_expand: + "(INF y:A. f y) = (\x. (INF y:A. f y x))" + by (simp add: fun_eq_iff INFI_apply) + +lemma SUPR_fun_expand: + "(SUP y:A. f y) = (\x. (SUP y:A. f y x))" + by (simp add: fun_eq_iff SUPR_apply) + lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]: assumes "\i. f i \ borel_measurable M" shows "(\x. (\\<^isub>\ i. f i x)) \ borel_measurable M" using assms unfolding psuminf_def - by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand]) + by (auto intro!: borel_measurable_SUP [unfolded SUPR_fun_expand]) + section "LIMSEQ is borel measurable" diff -r 051251fde456 -r a434f89a9962 src/HOL/Set.thy --- a/src/HOL/Set.thy Wed Dec 08 14:25:08 2010 +0100 +++ b/src/HOL/Set.thy Wed Dec 08 16:47:57 2010 +0100 @@ -533,6 +533,36 @@ by simp +subsubsection {* The empty set *} + +lemma empty_def: + "{} = {x. False}" + by (simp add: bot_fun_def bot_bool_def Collect_def) + +lemma empty_iff [simp]: "(c : {}) = False" + by (simp add: empty_def) + +lemma emptyE [elim!]: "a : {} ==> P" + by simp + +lemma empty_subsetI [iff]: "{} \ A" + -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *} + by blast + +lemma equals0I: "(!!y. y \ A ==> False) ==> A = {}" + by blast + +lemma equals0D: "A = {} ==> a \ A" + -- {* Use for reasoning about disjointness: @{text "A Int B = {}"} *} + by blast + +lemma ball_empty [simp]: "Ball {} P = True" + by (simp add: Ball_def) + +lemma bex_empty [simp]: "Bex {} P = False" + by (simp add: Bex_def) + + subsubsection {* The universal set -- UNIV *} abbreviation UNIV :: "'a set" where @@ -568,36 +598,6 @@ lemma UNIV_eq_I: "(\x. x \ A) \ UNIV = A" by auto - -subsubsection {* The empty set *} - -lemma empty_def: - "{} = {x. False}" - by (simp add: bot_fun_def bot_bool_def Collect_def) - -lemma empty_iff [simp]: "(c : {}) = False" - by (simp add: empty_def) - -lemma emptyE [elim!]: "a : {} ==> P" - by simp - -lemma empty_subsetI [iff]: "{} \ A" - -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *} - by blast - -lemma equals0I: "(!!y. y \ A ==> False) ==> A = {}" - by blast - -lemma equals0D: "A = {} ==> a \ A" - -- {* Use for reasoning about disjointness: @{text "A Int B = {}"} *} - by blast - -lemma ball_empty [simp]: "Ball {} P = True" - by (simp add: Ball_def) - -lemma bex_empty [simp]: "Bex {} P = False" - by (simp add: Bex_def) - lemma UNIV_not_empty [iff]: "UNIV ~= {}" by (blast elim: equalityE) @@ -647,7 +647,41 @@ lemma Compl_eq: "- A = {x. ~ x : A}" by blast -subsubsection {* Binary union -- Un *} +subsubsection {* Binary intersection *} + +abbreviation inter :: "'a set \ 'a set \ 'a set" (infixl "Int" 70) where + "op Int \ inf" + +notation (xsymbols) + inter (infixl "\" 70) + +notation (HTML output) + inter (infixl "\" 70) + +lemma Int_def: + "A \ B = {x. x \ A \ x \ B}" + by (simp add: inf_fun_def inf_bool_def Collect_def mem_def) + +lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)" + by (unfold Int_def) blast + +lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B" + by simp + +lemma IntD1: "c : A Int B ==> c:A" + by simp + +lemma IntD2: "c : A Int B ==> c:B" + by simp + +lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P" + by simp + +lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" + by (fact mono_inf) + + +subsubsection {* Binary union *} abbreviation union :: "'a set \ 'a set \ 'a set" (infixl "Un" 65) where "union \ sup" @@ -689,40 +723,6 @@ by (fact mono_sup) -subsubsection {* Binary intersection -- Int *} - -abbreviation inter :: "'a set \ 'a set \ 'a set" (infixl "Int" 70) where - "op Int \ inf" - -notation (xsymbols) - inter (infixl "\" 70) - -notation (HTML output) - inter (infixl "\" 70) - -lemma Int_def: - "A \ B = {x. x \ A \ x \ B}" - by (simp add: inf_fun_def inf_bool_def Collect_def mem_def) - -lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)" - by (unfold Int_def) blast - -lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B" - by simp - -lemma IntD1: "c : A Int B ==> c:A" - by simp - -lemma IntD2: "c : A Int B ==> c:B" - by simp - -lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P" - by simp - -lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" - by (fact mono_inf) - - subsubsection {* Set difference *} lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)"