# HG changeset patch # User haftmann # Date 1310595716 -7200 # Node ID 89082fd9e32de78b5ecc35a7ac2e0eab25ebcaa8 # Parent 93374f7448b6f9e91767c2cbb135c7dfcbd47a7e# Parent fcc5d3ffb6f572ec59d490fae19ee86d9d12b318 merged diff -r 93374f7448b6 -r 89082fd9e32d NEWS --- a/NEWS Wed Jul 13 22:16:19 2011 +0200 +++ b/NEWS Thu Jul 14 00:21:56 2011 +0200 @@ -60,6 +60,9 @@ *** HOL *** +* Classes bot and top require underlying partial order rather than preorder: +uniqueness of bot and top is guaranteed. INCOMPATIBILITY. + * Archimedian_Field.thy: floor now is defined as parameter of a separate type class floor_ceiling. diff -r 93374f7448b6 -r 89082fd9e32d src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Wed Jul 13 22:16:19 2011 +0200 +++ b/src/HOL/Complete_Lattice.thy Thu Jul 14 00:21:56 2011 +0200 @@ -392,15 +392,6 @@ lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" by (fact Inf_union_distrib) -lemma (in bounded_lattice_bot) bot_less: - -- {* FIXME: tighten classes bot, top to partial orders (uniqueness!), move lemmas there *} - "a \ bot \ bot < a" - by (auto simp add: less_le_not_le intro!: antisym) - -lemma (in bounded_lattice_top) less_top: - "a \ top \ a < top" - by (auto simp add: less_le_not_le intro!: antisym) - lemma (in complete_lattice) Inf_top_conv [no_atp]: "\A = \ \ (\x\A. x = \)" "\ = \A \ (\x\A. x = \)" @@ -478,21 +469,21 @@ "\(B`A) = (\x\A. B x)" by (rule sym) (fact INFI_def) -lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)" +lemma INT_iff [simp]: "b \ (\x\A. B x) \ (\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)" +lemma INT_I [intro!]: "(\x. x \ A \ b \ B x) \ b \ (\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" +lemma INT_D [elim, Pure.elim]: "b : (\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" +lemma INT_E [elim]: "b : (\x\A. B x) \ (b: B a \ R) \ (a~:A \ R) \ R" -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a:A"}. *} by (unfold INTER_def) blast lemma INT_cong [cong]: - "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)" + "A = B \ (\x. x:B \ C x = D x) \ (\x\A. C x) = (\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})" @@ -501,16 +492,16 @@ 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" +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)" +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)" +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)" @@ -523,7 +514,7 @@ by blast lemma INT_insert_distrib: - "u \ A ==> (\x\A. insert a (B x)) = insert a (\x\A. B x)" + "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)" @@ -534,23 +525,23 @@ 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)" + "(UNIV = (\x\A. B x)) = (\x\A. B x = UNIV)" + "((\x\A. B x) = UNIV) = (\x\A. B x = UNIV)" by blast+ -lemma INT_bool_eq: "(\b::bool. A b) = (A True \ A False)" +lemma INT_bool_eq: "(\b. 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) ==> + "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)" +lemma vimage_INT: "f -` (\x\A. B x) = (\x\A. f -` B x)" by blast @@ -583,40 +574,40 @@ by auto lemma UnionE [elim!]: - "A \ \C \ (\X. A\X \ X\C \ R) \ R" + "A \ \C \ (\X. A \ X \ X \ C \ R) \ R" by auto -lemma Union_upper: "B \ A ==> B \ Union A" +lemma Union_upper: "B \ A \ B \ \A" by (iprover intro: subsetI UnionI) -lemma Union_least: "(!!X. X \ A ==> X \ C) ==> Union A \ C" +lemma Union_least: "(\X. X \ A \ X \ C) \ \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({}) = {}" +lemma Union_empty [simp]: "\{} = {}" by blast -lemma Union_UNIV [simp]: "Union UNIV = UNIV" +lemma Union_UNIV [simp]: "\UNIV = UNIV" by blast -lemma Union_insert [simp]: "Union (insert a B) = a \ \B" +lemma Union_insert [simp]: "\insert a B = a \ \B" by blast -lemma Union_Un_distrib [simp]: "\(A Un B) = \A \ \B" +lemma Union_Un_distrib [simp]: "\(A \ B) = \A \ \B" by blast lemma Union_Int_subset: "\(A \ B) \ \A \ \B" by blast -lemma Union_empty_conv [simp,no_atp]: "(\A = {}) = (\x\A. x = {})" +lemma Union_empty_conv [simp,no_atp]: "(\A = {}) \ (\x\A. x = {})" by blast -lemma empty_Union_conv [simp,no_atp]: "({} = \A) = (\x\A. x = {})" +lemma empty_Union_conv [simp,no_atp]: "({} = \A) \ (\x\A. x = {})" by blast -lemma Union_disjoint: "(\C \ A = {}) = (\B\C. B \ A = {})" +lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" by blast lemma subset_Pow_Union: "A \ Pow (\A)" @@ -625,7 +616,7 @@ lemma Union_Pow_eq [simp]: "\(Pow A) = A" by blast -lemma Union_mono: "A \ B ==> \A \ \B" +lemma Union_mono: "A \ B \ \A \ \B" by blast @@ -666,7 +657,7 @@ *} -- {* to avoid eta-contraction of body *} lemma UNION_eq_Union_image: - "(\x\A. B x) = \(B`A)" + "(\x\A. B x) = \(B ` A)" by (fact SUPR_def) lemma Union_def: @@ -678,41 +669,41 @@ by (auto simp add: UNION_eq_Union_image Union_eq) lemma Union_image_eq [simp]: - "\(B`A) = (\x\A. B x)" + "\(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)" +lemma UN_iff [simp]: "(b: (\x\A. B x)) = (\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)" +lemma UN_I [intro]: "a:A \ b: B a \ b: (\x\A. B x)" -- {* The order of the premises presupposes that @{term A} is rigid; @{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" +lemma UN_E [elim!]: "b : (\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)" + "A = B \ (\x. x:B \ C x = D x) \ (\x\A. C x) = (\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)" + "A = B \ (\x. x:B =simp=> C x = D x) \ (\x\A. C x) = (\x\B. D x)" by (simp add: UNION_def simp_implies_def) -lemma image_eq_UN: "f`A = (UN x:A. {f x})" +lemma image_eq_UN: "f ` A = (\x\A. {f x})" by blast -lemma UN_upper: "a \ A ==> B a \ (\x\A. B x)" +lemma UN_upper: "a \ A \ B a \ (\x\A. B x)" by (fact le_SUPI) -lemma UN_least: "(!!x. x \ A ==> B x \ C) ==> (\x\A. B x) \ C" +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 [no_atp]: "{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)" +lemma UN_insert_distrib: "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" by blast lemma UN_empty [simp,no_atp]: "(\x\{}. B x) = {}" @@ -724,7 +715,7 @@ 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)" +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" @@ -749,8 +740,8 @@ 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 = {})" + "{} = (\x\A. B x) \ (\x\A. B x = {})" + "(\x\A. B x) = {} \ (\x\A. B x = {})" by blast+ lemma Collect_ex_eq [no_atp]: "{x. \y. P x y} = (\y. {x. P x y})" @@ -765,29 +756,29 @@ 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)" +lemma UN_bool_eq: "(\b. 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) ==> + "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)" +lemma vimage_Union: "f -` (\A) = (\X\A. f -` X)" by blast -lemma vimage_UN: "f-`(UN x:A. B x) = (UN x:A. f -` B x)" +lemma vimage_UN: "f -` (\x\A. B x) = (\x\A. f -` B x)" by blast -lemma vimage_eq_UN: "f-`B = (UN y: B. f-`{y})" +lemma vimage_eq_UN: "f -` B = (\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 +lemma image_UN: "f ` UNION A B = (\x\A. f ` B x)" + by blast subsection {* Distributive laws *} @@ -798,7 +789,7 @@ 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)" +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 @@ -810,7 +801,7 @@ 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)" +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)" @@ -833,10 +824,10 @@ subsection {* Complement *} -lemma Compl_UN [simp]: "-(\x\A. B x) = (\x\A. -B x)" +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)" +lemma Compl_INT [simp]: "- (\x\A. B x) = (\x\A. -B x)" by blast @@ -846,94 +837,85 @@ 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))" + "\a B C. (\x\C. insert a (B x)) = (if C={} then {} else insert a (\x\C. B x))" + "\A B C. (\x\C. A x Un B) = ((if C={} then {} else (\x\C. A x) Un B))" + "\A B C. (\x\C. A Un B x) = ((if C={} then {} else A Un (\x\C. B x)))" + "\A B C. (\x\C. A x Int B) = ((\x\C. A x) Int B)" + "\A B C. (\x\C. A Int B x) = (A Int (\x\C. B x))" + "\A B C. (\x\C. A x - B) = ((\x\C. A x) - B)" + "\A B C. (\x\C. A - B x) = (A - (\x\C. B x))" + "\A B. (UN x: \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))" + "\A B C. (\x\C. A x Int B) = (if C={} then UNIV else (\x\C. A x) Int B)" + "\A B C. (\x\C. A Int B x) = (if C={} then UNIV else A Int (\x\C. B x))" + "\A B C. (\x\C. A x - B) = (if C={} then UNIV else (\x\C. A x) - B)" + "\A B C. (\x\C. A - B x) = (if C={} then UNIV else A - (\x\C. B x))" + "\a B C. (\x\C. insert a (B x)) = insert a (\x\C. B x)" + "\A B C. (\x\C. A x Un B) = ((\x\C. A x) Un B)" + "\A B C. (\x\C. A Un B x) = (A Un (\x\C. B x))" + "\A B. (INT x: \A. B x) = (\y\A. INT x:y. B x)" + "\A B C. (INT z: UNION A B. C z) = (\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,no_atp]: - "!!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)" + "\A P Q. (\x\A. P x | Q) = ((\x\A. P x) | Q)" + "\A P Q. (\x\A. P | Q x) = (P | (\x\A. Q x))" + "\A P Q. (\x\A. P --> Q x) = (P --> (\x\A. Q x))" + "\A P Q. (\x\A. P x --> Q) = ((\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:\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) = (\x\A. P (f x))" + "\A P. (~(\x\A. P x)) = (\x\A. ~P x)" by auto lemma bex_simps [simp,no_atp]: - "!!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)" + "\A P Q. (\x\A. P x & Q) = ((\x\A. P x) & Q)" + "\A P Q. (\x\A. P & Q x) = (P & (\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:\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) = (\x\A. P (f x))" + "\A P. (~(\x\A. P x)) = (\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)" + "\a B C. insert a (\x\C. B x) = (if C={} then {a} else (\x\C. insert a (B x)))" + "\A B C. (\x\C. A x) Un B = (if C={} then B else (\x\C. A x Un B))" + "\A B C. A Un (\x\C. B x) = (if C={} then A else (\x\C. A Un B x))" + "\A B C. ((\x\C. A x) Int B) = (\x\C. A x Int B)" + "\A B C. (A Int (\x\C. B x)) = (\x\C. A Int B x)" + "\A B C. ((\x\C. A x) - B) = (\x\C. A x - B)" + "\A B C. (A - (\x\C. B x)) = (\x\C. A - B x)" + "\A B. (UN y:A. UN x:y. B x) = (UN x: \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)" + "\A B C. (\x\C. A x) Int B = (if C={} then B else (\x\C. A x Int B))" + "\A B C. A Int (\x\C. B x) = (if C={} then A else (\x\C. A Int B x))" + "\A B C. (\x\C. A x) - B = (if C={} then UNIV-B else (\x\C. A x - B))" + "\A B C. A - (\x\C. B x) = (if C={} then A else (\x\C. A - B x))" + "\a B C. insert a (\x\C. B x) = (\x\C. insert a (B x))" + "\A B C. ((\x\C. A x) Un B) = (\x\C. A x Un B)" + "\A B C. A Un (\x\C. B x) = (\x\C. A Un B x)" + "\A B. (\y\A. INT x:y. B x) = (INT x: \A. B x)" + "\A B C. (\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 diff -r 93374f7448b6 -r 89082fd9e32d src/HOL/Library/Option_ord.thy --- a/src/HOL/Library/Option_ord.thy Wed Jul 13 22:16:19 2011 +0200 +++ b/src/HOL/Library/Option_ord.thy Thu Jul 14 00:21:56 2011 +0200 @@ -58,7 +58,7 @@ instance option :: (linorder) linorder proof qed (auto simp add: less_eq_option_def less_option_def split: option.splits) -instantiation option :: (preorder) bot +instantiation option :: (order) bot begin definition "bot = None" diff -r 93374f7448b6 -r 89082fd9e32d src/HOL/Library/Quickcheck_Types.thy --- a/src/HOL/Library/Quickcheck_Types.thy Wed Jul 13 22:16:19 2011 +0200 +++ b/src/HOL/Library/Quickcheck_Types.thy Thu Jul 14 00:21:56 2011 +0200 @@ -108,7 +108,7 @@ instance bot :: (linorder) linorder proof qed (auto simp add: less_eq_bot_def less_bot_def split: bot.splits) -instantiation bot :: (preorder) bot +instantiation bot :: (order) bot begin definition "bot = Bot" @@ -208,7 +208,7 @@ instance top :: (linorder) linorder proof qed (auto simp add: less_eq_top_def less_top_def split: top.splits) -instantiation top :: (preorder) top +instantiation top :: (order) top begin definition "top = Top" diff -r 93374f7448b6 -r 89082fd9e32d src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Wed Jul 13 22:16:19 2011 +0200 +++ b/src/HOL/Orderings.thy Thu Jul 14 00:21:56 2011 +0200 @@ -1081,15 +1081,37 @@ done -subsection {* Top and bottom elements *} +subsection {* (Unique) top and bottom elements *} -class bot = preorder + +class bot = order + fixes bot :: 'a assumes bot_least [simp]: "bot \ x" +begin -class top = preorder + +lemma bot_unique: + "a \ bot \ a = bot" + by (auto simp add: intro: antisym) + +lemma bot_less: + "a \ bot \ bot < a" + by (auto simp add: less_le_not_le intro!: antisym) + +end + +class top = order + fixes top :: 'a assumes top_greatest [simp]: "x \ top" +begin + +lemma top_unique: + "top \ a \ a = top" + by (auto simp add: intro: antisym) + +lemma less_top: + "a \ top \ a < top" + by (auto simp add: less_le_not_le intro!: antisym) + +end subsection {* Dense orders *} diff -r 93374f7448b6 -r 89082fd9e32d src/HOL/Set.thy --- a/src/HOL/Set.thy Wed Jul 13 22:16:19 2011 +0200 +++ b/src/HOL/Set.thy Thu Jul 14 00:21:56 2011 +0200 @@ -416,6 +416,14 @@ lemma ball_one_point2 [simp]: "(ALL x:A. a = x --> P x) = (a:A --> P a)" by blast +lemma ball_conj_distrib: + "(\x\A. P x \ Q x) \ ((\x\A. P x) \ (\x\A. Q x))" + by blast + +lemma bex_disj_distrib: + "(\x\A. P x \ Q x) \ ((\x\A. P x) \ (\x\A. Q x))" + by blast + text {* Congruence rules *} @@ -535,7 +543,7 @@ lemma empty_def: "{} = {x. False}" - by (simp add: bot_fun_def bot_bool_def Collect_def) + by (simp add: bot_fun_def Collect_def) lemma empty_iff [simp]: "(c : {}) = False" by (simp add: empty_def) @@ -568,7 +576,7 @@ lemma UNIV_def: "UNIV = {x. True}" - by (simp add: top_fun_def top_bool_def Collect_def) + by (simp add: top_fun_def Collect_def) lemma UNIV_I [simp]: "x : UNIV" by (simp add: UNIV_def) @@ -627,7 +635,7 @@ subsubsection {* Set complement *} lemma Compl_iff [simp]: "(c \ -A) = (c \ A)" - by (simp add: mem_def fun_Compl_def bool_Compl_def) + by (simp add: mem_def fun_Compl_def) lemma ComplI [intro!]: "(c \ A ==> False) ==> c \ -A" by (unfold mem_def fun_Compl_def bool_Compl_def) blast @@ -638,7 +646,7 @@ right side of the notional turnstile ... *} lemma ComplD [dest!]: "c : -A ==> c~:A" - by (simp add: mem_def fun_Compl_def bool_Compl_def) + by (simp add: mem_def fun_Compl_def) lemmas ComplE = ComplD [elim_format] @@ -658,7 +666,7 @@ lemma Int_def: "A \ B = {x. x \ A \ x \ B}" - by (simp add: inf_fun_def inf_bool_def Collect_def mem_def) + by (simp add: inf_fun_def Collect_def mem_def) lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)" by (unfold Int_def) blast @@ -692,7 +700,7 @@ lemma Un_def: "A \ B = {x. x \ A \ x \ B}" - by (simp add: sup_fun_def sup_bool_def Collect_def mem_def) + by (simp add: sup_fun_def Collect_def mem_def) lemma Un_iff [simp]: "(c : A Un B) = (c:A | c:B)" by (unfold Un_def) blast @@ -724,7 +732,7 @@ subsubsection {* Set difference *} lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)" - by (simp add: mem_def fun_diff_def bool_diff_def) + by (simp add: mem_def fun_diff_def) lemma DiffI [intro!]: "c : A ==> c ~: B ==> c : A - B" by simp