# HG changeset patch # User haftmann # Date 1248278925 -7200 # Node ID 228905e0235057c08b5d71d415ab45d8abfec778 # Parent cabd6096892c67aa2bd5684426455979e1012861# Parent e271a64f03ff9913abd1bd7fa19ae8172f15f931 merged diff -r cabd6096892c -r 228905e02350 src/HOL/Complete_Lattice.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complete_Lattice.thy Wed Jul 22 18:08:45 2009 +0200 @@ -0,0 +1,794 @@ +(* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *) + +header {* Complete lattices, with special focus on sets *} + +theory Complete_Lattice +imports Set +begin + +notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + inf (infixl "\" 70) and + sup (infixl "\" 65) + + +subsection {* Abstract complete lattices *} + +class complete_lattice = lattice + bot + top + + fixes Inf :: "'a set \ 'a" ("\_" [900] 900) + and Sup :: "'a set \ 'a" ("\_" [900] 900) + assumes Inf_lower: "x \ A \ \A \ x" + and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" + assumes Sup_upper: "x \ A \ x \ \A" + and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" +begin + +lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" + by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) + +lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" + by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) + +lemma Inf_Univ: "\UNIV = \{}" + unfolding Sup_Inf by auto + +lemma Sup_Univ: "\UNIV = \{}" + unfolding Inf_Sup by auto + +lemma Inf_insert: "\insert a A = a \ \A" + by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) + +lemma Sup_insert: "\insert a A = a \ \A" + by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) + +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_insert_simp: + "\insert a A = (if A = {} then a else a \ \A)" + by (cases "A = {}") (simp_all, simp add: Inf_insert) + +lemma Sup_insert_simp: + "\insert a A = (if A = {} then a else a \ \A)" + by (cases "A = {}") (simp_all, simp add: Sup_insert) + +lemma Inf_binary: + "\{a, b} = a \ b" + by (auto simp add: Inf_insert_simp) + +lemma Sup_binary: + "\{a, b} = a \ b" + by (auto simp add: Sup_insert_simp) + +lemma bot_def: + "bot = \{}" + by (auto intro: antisym Sup_least) + +lemma top_def: + "top = \{}" + by (auto intro: antisym Inf_greatest) + +lemma sup_bot [simp]: + "x \ bot = x" + using bot_least [of x] by (simp add: le_iff_sup sup_commute) + +lemma inf_top [simp]: + "x \ top = x" + using top_greatest [of x] by (simp add: le_iff_inf inf_commute) + +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, 10] 10) + "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) + "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) + +translations + "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} "_SUP", +Syntax.preserve_binder_abs2_tr' @{const_syntax INFI} "_INF" +] *} -- {* to avoid eta-contraction of body *} + +context complete_lattice +begin + +lemma le_SUPI: "i : A \ M i \ (SUP i:A. M i)" + by (auto simp add: SUPR_def intro: Sup_upper) + +lemma SUP_leI: "(\i. i : A \ M i \ u) \ (SUP i:A. M i) \ u" + by (auto simp add: SUPR_def intro: Sup_least) + +lemma INF_leI: "i : A \ (INF i:A. M i) \ M i" + by (auto simp add: INFI_def intro: Inf_lower) + +lemma le_INFI: "(\i. i : A \ u \ M i) \ u \ (INF i:A. M i)" + by (auto simp add: INFI_def intro: Inf_greatest) + +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) + +end + + +subsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} + +instantiation bool :: complete_lattice +begin + +definition + Inf_bool_def: "\A \ (\x\A. x)" + +definition + Sup_bool_def: "\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: + "INFI = Ball" +proof (rule ext)+ + fix A :: "'a set" + fix P :: "'a \ bool" + show "(INF x:A. P x) \ (\x \ A. P x)" + by (auto simp add: Ball_def INFI_def Inf_bool_def) +qed + +lemma SUPR_bool_eq: + "SUPR = Bex" +proof (rule ext)+ + fix A :: "'a set" + fix P :: "'a \ bool" + show "(SUP x:A. P x) \ (\x \ A. P x)" + by (auto simp add: Bex_def SUPR_def Sup_bool_def) +qed + +instantiation "fun" :: (type, complete_lattice) complete_lattice +begin + +definition + Inf_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" + +definition + Sup_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" + +instance proof +qed (auto simp add: Inf_fun_def Sup_fun_def le_fun_def + intro: Inf_lower Sup_upper Inf_greatest Sup_least) + +end + +lemma Inf_empty_fun: + "\{} = (\_. \{})" + by (simp add: Inf_fun_def) + +lemma Sup_empty_fun: + "\{} = (\_. \{})" + by (simp add: Sup_fun_def) + + +subsection {* Union *} + +definition Union :: "'a set set \ 'a set" where + Sup_set_eq [symmetric]: "Union S = \S" + +notation (xsymbols) + Union ("\_" [90] 90) + +lemma Union_eq: + "\A = {x. \B \ A. x \ B}" +proof (rule set_ext) + 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: Sup_set_eq [symmetric] Sup_fun_def Sup_bool_def) (simp add: mem_def) +qed + +lemma Union_iff [simp, noatp]: + "A \ \C \ (\X\C. A\X)" + by (unfold Union_eq) blast + +lemma UnionI [intro]: + "X \ C \ A \ X \ A \ \C" + -- {* The order of the premises presupposes that @{term C} is rigid; + @{term A} may be flexible. *} + by auto + +lemma UnionE [elim!]: + "A \ \C \ (\X. A\X \ X\C \ R) \ R" + by auto + +lemma Union_upper: "B \ A ==> B \ Union A" + by (iprover intro: subsetI UnionI) + +lemma Union_least: "(!!X. X \ A ==> X \ C) ==> Union A \ C" + by (iprover intro: subsetI elim: UnionE dest: subsetD) + +lemma Un_eq_Union: "A \ B = \{A, B}" + by blast + +lemma Union_empty [simp]: "Union({}) = {}" + by blast + +lemma Union_UNIV [simp]: "Union UNIV = UNIV" + by blast + +lemma Union_insert [simp]: "Union (insert a B) = a \ \B" + by blast + +lemma Union_Un_distrib [simp]: "\(A Un B) = \A \ \B" + by blast + +lemma Union_Int_subset: "\(A \ B) \ \A \ \B" + by blast + +lemma Union_empty_conv [simp,noatp]: "(\A = {}) = (\x\A. x = {})" + by blast + +lemma empty_Union_conv [simp,noatp]: "({} = \A) = (\x\A. x = {})" + by blast + +lemma Union_disjoint: "(\C \ A = {}) = (\B\C. B \ A = {})" + by blast + +lemma subset_Pow_Union: "A \ Pow (\A)" + by blast + +lemma Union_Pow_eq [simp]: "\(Pow A) = A" + by blast + +lemma Union_mono: "A \ B ==> \A \ \B" + by blast + + +subsection {* Unions of families *} + +definition UNION :: "'a set \ ('a \ 'b set) \ 'b set" where + SUPR_set_eq [symmetric]: "UNION S f = (SUP x:S. f x)" + +syntax + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10) + +syntax (xsymbols) + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) + +syntax (latex output) + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) + +translations + "UN x y. B" == "UN x. UN y. B" + "UN x. B" == "CONST UNION CONST UNIV (%x. B)" + "UN x. B" == "UN x:CONST UNIV. B" + "UN x:A. B" == "CONST UNION A (%x. B)" + +text {* + Note the difference between ordinary xsymbol syntax of indexed + unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) + and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>1. B"}. The + former does not make the index expression a subscript of the + union/intersection symbol because this leads to problems with nested + subscripts in Proof General. +*} + +print_translation {* [ +Syntax.preserve_binder_abs2_tr' @{const_syntax UNION} "@UNION" +] *} -- {* to avoid eta-contraction of body *} + +lemma UNION_eq_Union_image: + "(\x\A. B x) = \(B`A)" + by (simp add: SUPR_def SUPR_set_eq [symmetric] Sup_set_eq) + +lemma Union_def: + "\S = (\x\S. x)" + by (simp add: UNION_eq_Union_image image_def) + +lemma UNION_def [noatp]: + "(\x\A. B x) = {y. \x\A. y \ B x}" + by (auto simp add: UNION_eq_Union_image Union_eq) + +lemma Union_image_eq [simp]: + "\(B`A) = (\x\A. B x)" + by (rule sym) (fact UNION_eq_Union_image) + +lemma UN_iff [simp]: "(b: (UN x:A. B x)) = (EX x:A. b: B x)" + by (unfold UNION_def) blast + +lemma UN_I [intro]: "a:A ==> b: B a ==> b: (UN x:A. B x)" + -- {* The order of the premises presupposes that @{term A} is rigid; + @{term b} may be flexible. *} + by auto + +lemma UN_E [elim!]: "b : (UN x:A. B x) ==> (!!x. x:A ==> b: B x ==> R) ==> R" + by (unfold UNION_def) blast + +lemma UN_cong [cong]: + "A = B ==> (!!x. x:B ==> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" + by (simp add: UNION_def) + +lemma strong_UN_cong: + "A = B ==> (!!x. x:B =simp=> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" + by (simp add: UNION_def simp_implies_def) + +lemma image_eq_UN: "f`A = (UN x:A. {f x})" + by blast + +lemma UN_upper: "a \ A ==> B a \ (\x\A. B x)" + by blast + +lemma UN_least: "(!!x. x \ A ==> B x \ C) ==> (\x\A. B x) \ C" + by (iprover intro: subsetI elim: UN_E dest: subsetD) + +lemma Collect_bex_eq [noatp]: "{x. \y\A. P x y} = (\y\A. {x. P x y})" + by blast + +lemma UN_insert_distrib: "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" + by blast + +lemma UN_empty [simp,noatp]: "(\x\{}. B x) = {}" + by blast + +lemma UN_empty2 [simp]: "(\x\A. {}) = {}" + by blast + +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 auto + +lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" + by blast + +lemma UN_Un[simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" + by blast + +lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" + by blast + +lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" + by blast + +lemma image_Union: "f ` \S = (\x\S. f ` x)" + by blast + +lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" + by auto + +lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by blast + +lemma UNION_empty_conv[simp]: + "({} = (UN x:A. B x)) = (\x\A. B x = {})" + "((UN x:A. B x) = {}) = (\x\A. B x = {})" +by blast+ + +lemma Collect_ex_eq [noatp]: "{x. \y. P x y} = (\y. {x. P x y})" + by blast + +lemma ball_UN: "(\z \ UNION A B. P z) = (\x\A. \z \ B x. P z)" + by blast + +lemma bex_UN: "(\z \ UNION A B. P z) = (\x\A. \z\B x. P z)" + by blast + +lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" + by (auto simp add: split_if_mem2) + +lemma UN_bool_eq: "(\b::bool. A b) = (A True \ A False)" + by (auto intro: bool_contrapos) + +lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" + by blast + +lemma UN_mono: + "A \ B ==> (!!x. x \ A ==> f x \ g x) ==> + (\x\A. f x) \ (\x\B. g x)" + by (blast dest: subsetD) + +lemma vimage_Union: "f -` (Union A) = (UN X:A. f -` X)" + by blast + +lemma vimage_UN: "f-`(UN x:A. B x) = (UN x:A. f -` B x)" + by blast + +lemma vimage_eq_UN: "f-`B = (UN y: B. f-`{y})" + -- {* NOT suitable for rewriting *} + by blast + +lemma image_UN: "(f ` (UNION A B)) = (UN x:A.(f ` (B x)))" +by blast + + +subsection {* Inter *} + +definition Inter :: "'a set set \ 'a set" where + Inf_set_eq [symmetric]: "Inter S = \S" + +notation (xsymbols) + Inter ("\_" [90] 90) + +lemma Inter_eq [code del]: + "\A = {x. \B \ A. x \ B}" +proof (rule set_ext) + 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 Inf_set_eq [symmetric]) (simp add: mem_def) +qed + +lemma Inter_iff [simp,noatp]: "(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]: "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,noatp]: + "(\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 *} + +definition INTER :: "'a set \ ('a \ 'b set) \ 'b set" where + INFI_set_eq [symmetric]: "INTER S f = (INF x:S. f x)" + +syntax + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [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, 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, 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} "@INTER" +] *} -- {* to avoid eta-contraction of body *} + +lemma INTER_eq_Inter_image: + "(\x\A. B x) = \(B`A)" + by (simp add: INFI_def INFI_set_eq [symmetric] Inf_set_eq) + +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]: "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 blast + +lemma INT_greatest: "(!!x. x \ A ==> C \ B x) ==> C \ (\x\A. B x)" + by (iprover intro: INT_I subsetI dest: subsetD) + +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 blast + +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)" + by blast + +lemma Int_Union2: "\B \ A = (\C\B. C \ A)" + by blast + +lemma Un_Union_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" + -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *} + -- {* Union of a family of unions *} + by blast + +lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" + -- {* Equivalent version *} + by blast + +lemma Un_Inter: "A \ \B = (\C\B. A \ C)" + by blast + +lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" + by blast + +lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" + -- {* Equivalent version *} + by blast + +lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" + -- {* Halmos, Naive Set Theory, page 35. *} + by blast + +lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" + by blast + +lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" + by blast + +lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" + by blast + + +subsection {* Complement *} + +lemma Compl_UN [simp]: "-(\x\A. B x) = (\x\A. -B x)" + by blast + +lemma Compl_INT [simp]: "-(\x\A. B x) = (\x\A. -B x)" + by blast + + +subsection {* Miniscoping and maxiscoping *} + +text {* \medskip Miniscoping: pushing in quantifiers and big Unions + and Intersections. *} + +lemma UN_simps [simp]: + "!!a B C. (UN x:C. insert a (B x)) = (if C={} then {} else insert a (UN x:C. B x))" + "!!A B C. (UN x:C. A x Un B) = ((if C={} then {} else (UN x:C. A x) Un B))" + "!!A B C. (UN x:C. A Un B x) = ((if C={} then {} else A Un (UN x:C. B x)))" + "!!A B C. (UN x:C. A x Int B) = ((UN x:C. A x) Int B)" + "!!A B C. (UN x:C. A Int B x) = (A Int (UN x:C. B x))" + "!!A B C. (UN x:C. A x - B) = ((UN x:C. A x) - B)" + "!!A B C. (UN x:C. A - B x) = (A - (INT x:C. B x))" + "!!A B. (UN x: Union A. B x) = (UN y:A. UN x:y. B x)" + "!!A B C. (UN z: UNION A B. C z) = (UN x:A. UN z: B(x). C z)" + "!!A B f. (UN x:f`A. B x) = (UN a:A. B (f a))" + by auto + +lemma INT_simps [simp]: + "!!A B C. (INT x:C. A x Int B) = (if C={} then UNIV else (INT x:C. A x) Int B)" + "!!A B C. (INT x:C. A Int B x) = (if C={} then UNIV else A Int (INT x:C. B x))" + "!!A B C. (INT x:C. A x - B) = (if C={} then UNIV else (INT x:C. A x) - B)" + "!!A B C. (INT x:C. A - B x) = (if C={} then UNIV else A - (UN x:C. B x))" + "!!a B C. (INT x:C. insert a (B x)) = insert a (INT x:C. B x)" + "!!A B C. (INT x:C. A x Un B) = ((INT x:C. A x) Un B)" + "!!A B C. (INT x:C. A Un B x) = (A Un (INT x:C. B x))" + "!!A B. (INT x: Union A. B x) = (INT y:A. INT x:y. B x)" + "!!A B C. (INT z: UNION A B. C z) = (INT x:A. INT z: B(x). C z)" + "!!A B f. (INT x:f`A. B x) = (INT a:A. B (f a))" + by auto + +lemma ball_simps [simp,noatp]: + "!!A P Q. (ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)" + "!!A P Q. (ALL x:A. P | Q x) = (P | (ALL x:A. Q x))" + "!!A P Q. (ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))" + "!!A P Q. (ALL x:A. P x --> Q) = ((EX x:A. P x) --> Q)" + "!!P. (ALL x:{}. P x) = True" + "!!P. (ALL x:UNIV. P x) = (ALL x. P x)" + "!!a B P. (ALL x:insert a B. P x) = (P a & (ALL x:B. P x))" + "!!A P. (ALL x:Union A. P x) = (ALL y:A. ALL x:y. P x)" + "!!A B P. (ALL x: UNION A B. P x) = (ALL a:A. ALL x: B a. P x)" + "!!P Q. (ALL x:Collect Q. P x) = (ALL x. Q x --> P x)" + "!!A P f. (ALL x:f`A. P x) = (ALL x:A. P (f x))" + "!!A P. (~(ALL x:A. P x)) = (EX x:A. ~P x)" + by auto + +lemma bex_simps [simp,noatp]: + "!!A P Q. (EX x:A. P x & Q) = ((EX x:A. P x) & Q)" + "!!A P Q. (EX x:A. P & Q x) = (P & (EX x:A. Q x))" + "!!P. (EX x:{}. P x) = False" + "!!P. (EX x:UNIV. P x) = (EX x. P x)" + "!!a B P. (EX x:insert a B. P x) = (P(a) | (EX x:B. P x))" + "!!A P. (EX x:Union A. P x) = (EX y:A. EX x:y. P x)" + "!!A B P. (EX x: UNION A B. P x) = (EX a:A. EX x:B a. P x)" + "!!P Q. (EX x:Collect Q. P x) = (EX x. Q x & P x)" + "!!A P f. (EX x:f`A. P x) = (EX x:A. P (f x))" + "!!A P. (~(EX x:A. P x)) = (ALL x:A. ~P x)" + by auto + +lemma ball_conj_distrib: + "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))" + by blast + +lemma bex_disj_distrib: + "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))" + by blast + + +text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} + +lemma UN_extend_simps: + "!!a B C. insert a (UN x:C. B x) = (if C={} then {a} else (UN x:C. insert a (B x)))" + "!!A B C. (UN x:C. A x) Un B = (if C={} then B else (UN x:C. A x Un B))" + "!!A B C. A Un (UN x:C. B x) = (if C={} then A else (UN x:C. A Un B x))" + "!!A B C. ((UN x:C. A x) Int B) = (UN x:C. A x Int B)" + "!!A B C. (A Int (UN x:C. B x)) = (UN x:C. A Int B x)" + "!!A B C. ((UN x:C. A x) - B) = (UN x:C. A x - B)" + "!!A B C. (A - (INT x:C. B x)) = (UN x:C. A - B x)" + "!!A B. (UN y:A. UN x:y. B x) = (UN x: Union A. B x)" + "!!A B C. (UN x:A. UN z: B(x). C z) = (UN z: UNION A B. C z)" + "!!A B f. (UN a:A. B (f a)) = (UN x:f`A. B x)" + by auto + +lemma INT_extend_simps: + "!!A B C. (INT x:C. A x) Int B = (if C={} then B else (INT x:C. A x Int B))" + "!!A B C. A Int (INT x:C. B x) = (if C={} then A else (INT x:C. A Int B x))" + "!!A B C. (INT x:C. A x) - B = (if C={} then UNIV-B else (INT x:C. A x - B))" + "!!A B C. A - (UN x:C. B x) = (if C={} then A else (INT x:C. A - B x))" + "!!a B C. insert a (INT x:C. B x) = (INT x:C. insert a (B x))" + "!!A B C. ((INT x:C. A x) Un B) = (INT x:C. A x Un B)" + "!!A B C. A Un (INT x:C. B x) = (INT x:C. A Un B x)" + "!!A B. (INT y:A. INT x:y. B x) = (INT x: Union A. B x)" + "!!A B C. (INT x:A. INT z: B(x). C z) = (INT z: UNION A B. C z)" + "!!A B f. (INT a:A. B (f a)) = (INT x:f`A. B x)" + by auto + + +no_notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + inf (infixl "\" 70) and + sup (infixl "\" 65) and + Inf ("\_" [900] 900) and + Sup ("\_" [900] 900) + +lemmas mem_simps = + insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff + mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff + -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} + +end diff -r cabd6096892c -r 228905e02350 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/Fun.thy Wed Jul 22 18:08:45 2009 +0200 @@ -6,7 +6,7 @@ header {* Notions about functions *} theory Fun -imports Set +imports Complete_Lattice begin text{*As a simplification rule, it replaces all function equalities by diff -r cabd6096892c -r 228905e02350 src/HOL/HoareParallel/OG_Hoare.thy --- a/src/HOL/HoareParallel/OG_Hoare.thy Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/HoareParallel/OG_Hoare.thy Wed Jul 22 18:08:45 2009 +0200 @@ -441,7 +441,7 @@ apply clarify apply(frule Parallel_length_post_PStar) apply clarify - apply(drule_tac j=xa in Parallel_Strong_Soundness) + apply(drule_tac j=xb in Parallel_Strong_Soundness) apply clarify apply simp apply force diff -r cabd6096892c -r 228905e02350 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/IsaMakefile Wed Jul 22 18:08:45 2009 +0200 @@ -117,6 +117,7 @@ @$(ISABELLE_TOOL) usedir -b -f base.ML -d false -g false $(OUT)/Pure HOL-Base PLAIN_DEPENDENCIES = $(BASE_DEPENDENCIES)\ + Complete_Lattice.thy \ Datatype.thy \ Divides.thy \ Extraction.thy \ diff -r cabd6096892c -r 228905e02350 src/HOL/Library/Executable_Set.thy --- a/src/HOL/Library/Executable_Set.thy Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/Library/Executable_Set.thy Wed Jul 22 18:08:45 2009 +0200 @@ -75,8 +75,8 @@ "op \" ("{*Fset.union*}") "op \" ("{*Fset.inter*}") "op - \ 'a set \ 'a set \ 'a set" ("{*flip Fset.subtract*}") - "Set.Union" ("{*Fset.Union*}") - "Set.Inter" ("{*Fset.Inter*}") + "Complete_Lattice.Union" ("{*Fset.Union*}") + "Complete_Lattice.Inter" ("{*Fset.Inter*}") card ("{*Fset.card*}") fold ("{*foldl o flip*}") diff -r cabd6096892c -r 228905e02350 src/HOL/Library/Fset.thy --- a/src/HOL/Library/Fset.thy Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/Library/Fset.thy Wed Jul 22 18:08:45 2009 +0200 @@ -160,7 +160,7 @@ qed definition Inter :: "'a fset fset \ 'a fset" where - [simp]: "Inter A = Fset (Set.Inter (member ` member A))" + [simp]: "Inter A = Fset (Complete_Lattice.Inter (member ` member A))" lemma Inter_inter [code]: "Inter (Set (A # As)) = foldl inter A As" @@ -174,7 +174,7 @@ qed definition Union :: "'a fset fset \ 'a fset" where - [simp]: "Union A = Fset (Set.Union (member ` member A))" + [simp]: "Union A = Fset (Complete_Lattice.Union (member ` member A))" lemma Union_union [code]: "Union (Set As) = foldl union empty As" diff -r cabd6096892c -r 228905e02350 src/HOL/Library/Lattice_Syntax.thy --- a/src/HOL/Library/Lattice_Syntax.thy Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/Library/Lattice_Syntax.thy Wed Jul 22 18:08:45 2009 +0200 @@ -4,16 +4,16 @@ (*<*) theory Lattice_Syntax -imports Set +imports Complete_Lattice begin notation + top ("\") and + bot ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) and Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) and - top ("\") and - bot ("\") + Sup ("\_" [900] 900) end (*>*) \ No newline at end of file diff -r cabd6096892c -r 228905e02350 src/HOL/Set.thy --- a/src/HOL/Set.thy Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/Set.thy Wed Jul 22 18:08:45 2009 +0200 @@ -1,6 +1,4 @@ -(* Title: HOL/Set.thy - Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel -*) +(* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel *) header {* Set theory for higher-order logic *} @@ -1733,789 +1731,4 @@ val vimage_Un = @{thm vimage_Un} *} - -subsection {* Complete lattices *} - -notation - less_eq (infix "\" 50) and - less (infix "\" 50) and - inf (infixl "\" 70) and - sup (infixl "\" 65) - -class complete_lattice = lattice + bot + top + - fixes Inf :: "'a set \ 'a" ("\_" [900] 900) - and Sup :: "'a set \ 'a" ("\_" [900] 900) - assumes Inf_lower: "x \ A \ \A \ x" - and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" - assumes Sup_upper: "x \ A \ x \ \A" - and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" -begin - -lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Inf_Univ: "\UNIV = \{}" - unfolding Sup_Inf by auto - -lemma Sup_Univ: "\UNIV = \{}" - unfolding Inf_Sup by auto - -lemma Inf_insert: "\insert a A = a \ \A" - by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) - -lemma Sup_insert: "\insert a A = a \ \A" - by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) - -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_insert_simp: - "\insert a A = (if A = {} then a else a \ \A)" - by (cases "A = {}") (simp_all, simp add: Inf_insert) - -lemma Sup_insert_simp: - "\insert a A = (if A = {} then a else a \ \A)" - by (cases "A = {}") (simp_all, simp add: Sup_insert) - -lemma Inf_binary: - "\{a, b} = a \ b" - by (auto simp add: Inf_insert_simp) - -lemma Sup_binary: - "\{a, b} = a \ b" - by (auto simp add: Sup_insert_simp) - -lemma bot_def: - "bot = \{}" - by (auto intro: antisym Sup_least) - -lemma top_def: - "top = \{}" - by (auto intro: antisym Inf_greatest) - -lemma sup_bot [simp]: - "x \ bot = x" - using bot_least [of x] by (simp add: le_iff_sup sup_commute) - -lemma inf_top [simp]: - "x \ top = x" - using top_greatest [of x] by (simp add: le_iff_inf inf_commute) - -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, 10] 10) - "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) - "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) - -translations - "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} "_SUP", -Syntax.preserve_binder_abs2_tr' @{const_syntax INFI} "_INF" -] *} -- {* to avoid eta-contraction of body *} - -context complete_lattice -begin - -lemma le_SUPI: "i : A \ M i \ (SUP i:A. M i)" - by (auto simp add: SUPR_def intro: Sup_upper) - -lemma SUP_leI: "(\i. i : A \ M i \ u) \ (SUP i:A. M i) \ u" - by (auto simp add: SUPR_def intro: Sup_least) - -lemma INF_leI: "i : A \ (INF i:A. M i) \ M i" - by (auto simp add: INFI_def intro: Inf_lower) - -lemma le_INFI: "(\i. i : A \ u \ M i) \ u \ (INF i:A. M i)" - by (auto simp add: INFI_def intro: Inf_greatest) - -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) - -end - - -subsubsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} - -instantiation bool :: complete_lattice -begin - -definition - Inf_bool_def: "\A \ (\x\A. x)" - -definition - Sup_bool_def: "\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: - "INFI = Ball" -proof (rule ext)+ - fix A :: "'a set" - fix P :: "'a \ bool" - show "(INF x:A. P x) \ (\x \ A. P x)" - by (auto simp add: Ball_def INFI_def Inf_bool_def) -qed - -lemma SUPR_bool_eq: - "SUPR = Bex" -proof (rule ext)+ - fix A :: "'a set" - fix P :: "'a \ bool" - show "(SUP x:A. P x) \ (\x \ A. P x)" - by (auto simp add: Bex_def SUPR_def Sup_bool_def) -qed - -instantiation "fun" :: (type, complete_lattice) complete_lattice -begin - -definition - Inf_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" - -definition - Sup_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" - -instance proof -qed (auto simp add: Inf_fun_def Sup_fun_def le_fun_def - intro: Inf_lower Sup_upper Inf_greatest Sup_least) - -end - -lemma Inf_empty_fun: - "\{} = (\_. \{})" - by (simp add: Inf_fun_def) - -lemma Sup_empty_fun: - "\{} = (\_. \{})" - by (simp add: Sup_fun_def) - - -subsubsection {* Union *} - -definition Union :: "'a set set \ 'a set" where - Sup_set_eq [symmetric]: "Union S = \S" - -notation (xsymbols) - Union ("\_" [90] 90) - -lemma Union_eq: - "\A = {x. \B \ A. x \ B}" -proof (rule set_ext) - 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: Sup_set_eq [symmetric] Sup_fun_def Sup_bool_def) (simp add: mem_def) -qed - -lemma Union_iff [simp, noatp]: - "A \ \C \ (\X\C. A\X)" - by (unfold Union_eq) blast - -lemma UnionI [intro]: - "X \ C \ A \ X \ A \ \C" - -- {* The order of the premises presupposes that @{term C} is rigid; - @{term A} may be flexible. *} - by auto - -lemma UnionE [elim!]: - "A \ \C \ (\X. A\X \ X\C \ R) \ R" - by auto - -lemma Union_upper: "B \ A ==> B \ Union A" - by (iprover intro: subsetI UnionI) - -lemma Union_least: "(!!X. X \ A ==> X \ C) ==> Union A \ C" - by (iprover intro: subsetI elim: UnionE dest: subsetD) - -lemma Un_eq_Union: "A \ B = \{A, B}" - by blast - -lemma Union_empty [simp]: "Union({}) = {}" - by blast - -lemma Union_UNIV [simp]: "Union UNIV = UNIV" - by blast - -lemma Union_insert [simp]: "Union (insert a B) = a \ \B" - by blast - -lemma Union_Un_distrib [simp]: "\(A Un B) = \A \ \B" - by blast - -lemma Union_Int_subset: "\(A \ B) \ \A \ \B" - by blast - -lemma Union_empty_conv [simp,noatp]: "(\A = {}) = (\x\A. x = {})" - by blast - -lemma empty_Union_conv [simp,noatp]: "({} = \A) = (\x\A. x = {})" - by blast - -lemma Union_disjoint: "(\C \ A = {}) = (\B\C. B \ A = {})" - by blast - -lemma subset_Pow_Union: "A \ Pow (\A)" - by blast - -lemma Union_Pow_eq [simp]: "\(Pow A) = A" - by blast - -lemma Union_mono: "A \ B ==> \A \ \B" - by blast - - -subsubsection {* Unions of families *} - -definition UNION :: "'a set \ ('a \ 'b set) \ 'b set" where - SUPR_set_eq [symmetric]: "UNION S f = (SUP x:S. f x)" - -syntax - "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) - "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10) - -syntax (xsymbols) - "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) - -syntax (latex output) - "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) - "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) - -translations - "UN x y. B" == "UN x. UN y. B" - "UN x. B" == "CONST UNION CONST UNIV (%x. B)" - "UN x. B" == "UN x:CONST UNIV. B" - "UN x:A. B" == "CONST UNION A (%x. B)" - -text {* - Note the difference between ordinary xsymbol syntax of indexed - unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) - and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>1. B"}. The - former does not make the index expression a subscript of the - union/intersection symbol because this leads to problems with nested - subscripts in Proof General. -*} - -print_translation {* [ -Syntax.preserve_binder_abs2_tr' @{const_syntax UNION} "@UNION" -] *} -- {* to avoid eta-contraction of body *} - -lemma UNION_eq_Union_image: - "(\x\A. B x) = \(B`A)" - by (simp add: SUPR_def SUPR_set_eq [symmetric] Sup_set_eq) - -lemma Union_def: - "\S = (\x\S. x)" - by (simp add: UNION_eq_Union_image image_def) - -lemma UNION_def [noatp]: - "(\x\A. B x) = {y. \x\A. y \ B x}" - by (auto simp add: UNION_eq_Union_image Union_eq) - -lemma Union_image_eq [simp]: - "\(B`A) = (\x\A. B x)" - by (rule sym) (fact UNION_eq_Union_image) - -lemma UN_iff [simp]: "(b: (UN x:A. B x)) = (EX x:A. b: B x)" - by (unfold UNION_def) blast - -lemma UN_I [intro]: "a:A ==> b: B a ==> b: (UN x:A. B x)" - -- {* The order of the premises presupposes that @{term A} is rigid; - @{term b} may be flexible. *} - by auto - -lemma UN_E [elim!]: "b : (UN x:A. B x) ==> (!!x. x:A ==> b: B x ==> R) ==> R" - by (unfold UNION_def) blast - -lemma UN_cong [cong]: - "A = B ==> (!!x. x:B ==> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" - by (simp add: UNION_def) - -lemma strong_UN_cong: - "A = B ==> (!!x. x:B =simp=> C x = D x) ==> (UN x:A. C x) = (UN x:B. D x)" - by (simp add: UNION_def simp_implies_def) - -lemma image_eq_UN: "f`A = (UN x:A. {f x})" - by blast - -lemma UN_upper: "a \ A ==> B a \ (\x\A. B x)" - by blast - -lemma UN_least: "(!!x. x \ A ==> B x \ C) ==> (\x\A. B x) \ C" - by (iprover intro: subsetI elim: UN_E dest: subsetD) - -lemma Collect_bex_eq [noatp]: "{x. \y\A. P x y} = (\y\A. {x. P x y})" - by blast - -lemma UN_insert_distrib: "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" - by blast - -lemma UN_empty [simp,noatp]: "(\x\{}. B x) = {}" - by blast - -lemma UN_empty2 [simp]: "(\x\A. {}) = {}" - by blast - -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 auto - -lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" - by blast - -lemma UN_Un[simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" - by blast - -lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" - by blast - -lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" - by blast - -lemma image_Union: "f ` \S = (\x\S. f ` x)" - by blast - -lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" - by auto - -lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by blast - -lemma UNION_empty_conv[simp]: - "({} = (UN x:A. B x)) = (\x\A. B x = {})" - "((UN x:A. B x) = {}) = (\x\A. B x = {})" -by blast+ - -lemma Collect_ex_eq [noatp]: "{x. \y. P x y} = (\y. {x. P x y})" - by blast - -lemma ball_UN: "(\z \ UNION A B. P z) = (\x\A. \z \ B x. P z)" - by blast - -lemma bex_UN: "(\z \ UNION A B. P z) = (\x\A. \z\B x. P z)" - by blast - -lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" - by (auto simp add: split_if_mem2) - -lemma UN_bool_eq: "(\b::bool. A b) = (A True \ A False)" - by (auto intro: bool_contrapos) - -lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" - by blast - -lemma UN_mono: - "A \ B ==> (!!x. x \ A ==> f x \ g x) ==> - (\x\A. f x) \ (\x\B. g x)" - by (blast dest: subsetD) - -lemma vimage_Union: "f -` (Union A) = (UN X:A. f -` X)" - by blast - -lemma vimage_UN: "f-`(UN x:A. B x) = (UN x:A. f -` B x)" - by blast - -lemma vimage_eq_UN: "f-`B = (UN y: B. f-`{y})" - -- {* NOT suitable for rewriting *} - by blast - -lemma image_UN: "(f ` (UNION A B)) = (UN x:A.(f ` (B x)))" -by blast - - -subsubsection {* Inter *} - -definition Inter :: "'a set set \ 'a set" where - Inf_set_eq [symmetric]: "Inter S = \S" - -notation (xsymbols) - Inter ("\_" [90] 90) - -lemma Inter_eq [code del]: - "\A = {x. \B \ A. x \ B}" -proof (rule set_ext) - 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 Inf_set_eq [symmetric]) (simp add: mem_def) -qed - -lemma Inter_iff [simp,noatp]: "(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]: "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,noatp]: - "(\A = UNIV) = (\x\A. x = UNIV)" - "(UNIV = \A) = (\x\A. x = UNIV)" - by blast+ - -lemma Inter_anti_mono: "B \ A ==> \A \ \B" - by blast - - -subsubsection {* Intersections of families *} - -definition INTER :: "'a set \ ('a \ 'b set) \ 'b set" where - INFI_set_eq [symmetric]: "INTER S f = (INF x:S. f x)" - -syntax - "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) - "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [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, 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, 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} "@INTER" -] *} -- {* to avoid eta-contraction of body *} - -lemma INTER_eq_Inter_image: - "(\x\A. B x) = \(B`A)" - by (simp add: INFI_def INFI_set_eq [symmetric] Inf_set_eq) - -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]: "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 blast - -lemma INT_greatest: "(!!x. x \ A ==> C \ B x) ==> C \ (\x\A. B x)" - by (iprover intro: INT_I subsetI dest: subsetD) - -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 blast - -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 - - -subsubsection {* Distributive laws *} - -lemma Int_Union: "A \ \B = (\C\B. A \ C)" - by blast - -lemma Int_Union2: "\B \ A = (\C\B. C \ A)" - by blast - -lemma Un_Union_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" - -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *} - -- {* Union of a family of unions *} - by blast - -lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" - -- {* Equivalent version *} - by blast - -lemma Un_Inter: "A \ \B = (\C\B. A \ C)" - by blast - -lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A`C) \ \(B`C)" - by blast - -lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" - -- {* Equivalent version *} - by blast - -lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" - -- {* Halmos, Naive Set Theory, page 35. *} - by blast - -lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" - by blast - -lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" - by blast - -lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" - by blast - - -subsubsection {* Complement *} - -lemma Compl_UN [simp]: "-(\x\A. B x) = (\x\A. -B x)" - by blast - -lemma Compl_INT [simp]: "-(\x\A. B x) = (\x\A. -B x)" - by blast - - -subsubsection {* Miniscoping and maxiscoping *} - -text {* \medskip Miniscoping: pushing in quantifiers and big Unions - and Intersections. *} - -lemma UN_simps [simp]: - "!!a B C. (UN x:C. insert a (B x)) = (if C={} then {} else insert a (UN x:C. B x))" - "!!A B C. (UN x:C. A x Un B) = ((if C={} then {} else (UN x:C. A x) Un B))" - "!!A B C. (UN x:C. A Un B x) = ((if C={} then {} else A Un (UN x:C. B x)))" - "!!A B C. (UN x:C. A x Int B) = ((UN x:C. A x) Int B)" - "!!A B C. (UN x:C. A Int B x) = (A Int (UN x:C. B x))" - "!!A B C. (UN x:C. A x - B) = ((UN x:C. A x) - B)" - "!!A B C. (UN x:C. A - B x) = (A - (INT x:C. B x))" - "!!A B. (UN x: Union A. B x) = (UN y:A. UN x:y. B x)" - "!!A B C. (UN z: UNION A B. C z) = (UN x:A. UN z: B(x). C z)" - "!!A B f. (UN x:f`A. B x) = (UN a:A. B (f a))" - by auto - -lemma INT_simps [simp]: - "!!A B C. (INT x:C. A x Int B) = (if C={} then UNIV else (INT x:C. A x) Int B)" - "!!A B C. (INT x:C. A Int B x) = (if C={} then UNIV else A Int (INT x:C. B x))" - "!!A B C. (INT x:C. A x - B) = (if C={} then UNIV else (INT x:C. A x) - B)" - "!!A B C. (INT x:C. A - B x) = (if C={} then UNIV else A - (UN x:C. B x))" - "!!a B C. (INT x:C. insert a (B x)) = insert a (INT x:C. B x)" - "!!A B C. (INT x:C. A x Un B) = ((INT x:C. A x) Un B)" - "!!A B C. (INT x:C. A Un B x) = (A Un (INT x:C. B x))" - "!!A B. (INT x: Union A. B x) = (INT y:A. INT x:y. B x)" - "!!A B C. (INT z: UNION A B. C z) = (INT x:A. INT z: B(x). C z)" - "!!A B f. (INT x:f`A. B x) = (INT a:A. B (f a))" - by auto - -lemma ball_simps [simp,noatp]: - "!!A P Q. (ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)" - "!!A P Q. (ALL x:A. P | Q x) = (P | (ALL x:A. Q x))" - "!!A P Q. (ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))" - "!!A P Q. (ALL x:A. P x --> Q) = ((EX x:A. P x) --> Q)" - "!!P. (ALL x:{}. P x) = True" - "!!P. (ALL x:UNIV. P x) = (ALL x. P x)" - "!!a B P. (ALL x:insert a B. P x) = (P a & (ALL x:B. P x))" - "!!A P. (ALL x:Union A. P x) = (ALL y:A. ALL x:y. P x)" - "!!A B P. (ALL x: UNION A B. P x) = (ALL a:A. ALL x: B a. P x)" - "!!P Q. (ALL x:Collect Q. P x) = (ALL x. Q x --> P x)" - "!!A P f. (ALL x:f`A. P x) = (ALL x:A. P (f x))" - "!!A P. (~(ALL x:A. P x)) = (EX x:A. ~P x)" - by auto - -lemma bex_simps [simp,noatp]: - "!!A P Q. (EX x:A. P x & Q) = ((EX x:A. P x) & Q)" - "!!A P Q. (EX x:A. P & Q x) = (P & (EX x:A. Q x))" - "!!P. (EX x:{}. P x) = False" - "!!P. (EX x:UNIV. P x) = (EX x. P x)" - "!!a B P. (EX x:insert a B. P x) = (P(a) | (EX x:B. P x))" - "!!A P. (EX x:Union A. P x) = (EX y:A. EX x:y. P x)" - "!!A B P. (EX x: UNION A B. P x) = (EX a:A. EX x:B a. P x)" - "!!P Q. (EX x:Collect Q. P x) = (EX x. Q x & P x)" - "!!A P f. (EX x:f`A. P x) = (EX x:A. P (f x))" - "!!A P. (~(EX x:A. P x)) = (ALL x:A. ~P x)" - by auto - -lemma ball_conj_distrib: - "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))" - by blast - -lemma bex_disj_distrib: - "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))" - by blast - - -text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} - -lemma UN_extend_simps: - "!!a B C. insert a (UN x:C. B x) = (if C={} then {a} else (UN x:C. insert a (B x)))" - "!!A B C. (UN x:C. A x) Un B = (if C={} then B else (UN x:C. A x Un B))" - "!!A B C. A Un (UN x:C. B x) = (if C={} then A else (UN x:C. A Un B x))" - "!!A B C. ((UN x:C. A x) Int B) = (UN x:C. A x Int B)" - "!!A B C. (A Int (UN x:C. B x)) = (UN x:C. A Int B x)" - "!!A B C. ((UN x:C. A x) - B) = (UN x:C. A x - B)" - "!!A B C. (A - (INT x:C. B x)) = (UN x:C. A - B x)" - "!!A B. (UN y:A. UN x:y. B x) = (UN x: Union A. B x)" - "!!A B C. (UN x:A. UN z: B(x). C z) = (UN z: UNION A B. C z)" - "!!A B f. (UN a:A. B (f a)) = (UN x:f`A. B x)" - by auto - -lemma INT_extend_simps: - "!!A B C. (INT x:C. A x) Int B = (if C={} then B else (INT x:C. A x Int B))" - "!!A B C. A Int (INT x:C. B x) = (if C={} then A else (INT x:C. A Int B x))" - "!!A B C. (INT x:C. A x) - B = (if C={} then UNIV-B else (INT x:C. A x - B))" - "!!A B C. A - (UN x:C. B x) = (if C={} then A else (INT x:C. A - B x))" - "!!a B C. insert a (INT x:C. B x) = (INT x:C. insert a (B x))" - "!!A B C. ((INT x:C. A x) Un B) = (INT x:C. A x Un B)" - "!!A B C. A Un (INT x:C. B x) = (INT x:C. A Un B x)" - "!!A B. (INT y:A. INT x:y. B x) = (INT x: Union A. B x)" - "!!A B C. (INT x:A. INT z: B(x). C z) = (INT z: UNION A B. C z)" - "!!A B f. (INT a:A. B (f a)) = (INT x:f`A. B x)" - by auto - - -no_notation - less_eq (infix "\" 50) and - less (infix "\" 50) and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) - -lemmas mem_simps = - insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff - mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff - -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} - -end diff -r cabd6096892c -r 228905e02350 src/HOL/UNITY/Comp/Alloc.thy --- a/src/HOL/UNITY/Comp/Alloc.thy Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/UNITY/Comp/Alloc.thy Wed Jul 22 18:08:45 2009 +0200 @@ -1021,7 +1021,7 @@ LeadsTo {s. h pfixLe (sub i o allocGiv) s})" apply (simp only: o_apply sub_def) apply (insert Alloc_Progress [THEN rename_guarantees_sysOfAlloc_I]) - apply (simp add: o_def del: Set.INT_iff); + apply (simp add: o_def del: INT_iff) apply (erule component_guaranteesD) apply (auto simp add: System_Increasing_allocRel [simplified sub_apply o_def] diff -r cabd6096892c -r 228905e02350 src/HOL/UNITY/ProgressSets.thy --- a/src/HOL/UNITY/ProgressSets.thy Wed Jul 22 15:28:49 2009 +0200 +++ b/src/HOL/UNITY/ProgressSets.thy Wed Jul 22 18:08:45 2009 +0200 @@ -44,7 +44,7 @@ lemma UN_in_lattice: "[|lattice L; !!i. i\I ==> r i \ L|] ==> (\i\I. r i) \ L" -apply (simp add: Set.UN_eq) +apply (simp add: UN_eq) apply (blast intro: Union_in_lattice) done