bot comes before top, inf before sup etc.
authorhaftmann
Wed, 08 Dec 2010 15:05:46 +0100
changeset 413309ff94e7cc3b3
parent 41329 fb1e5377143d
child 41331 c987429a1298
bot comes before top, inf before sup etc.
src/HOL/Complete_Lattice.thy
src/HOL/Lattices.thy
src/HOL/Library/Lattice_Syntax.thy
src/HOL/Orderings.thy
src/HOL/Predicate.thy
src/HOL/Set.thy
     1.1 --- a/src/HOL/Complete_Lattice.thy	Wed Dec 08 14:52:23 2010 +0100
     1.2 +++ b/src/HOL/Complete_Lattice.thy	Wed Dec 08 15:05:46 2010 +0100
     1.3 @@ -82,11 +82,21 @@
     1.4    "\<Squnion>{a, b} = a \<squnion> b"
     1.5    by (simp add: Sup_empty Sup_insert)
     1.6  
     1.7 +lemma le_Inf_iff: "b \<sqsubseteq> Inf A \<longleftrightarrow> (\<forall>a\<in>A. b \<sqsubseteq> a)"
     1.8 +  by (auto intro: Inf_greatest dest: Inf_lower)
     1.9 +
    1.10  lemma Sup_le_iff: "Sup A \<sqsubseteq> b \<longleftrightarrow> (\<forall>a\<in>A. a \<sqsubseteq> b)"
    1.11    by (auto intro: Sup_least dest: Sup_upper)
    1.12  
    1.13 -lemma le_Inf_iff: "b \<sqsubseteq> Inf A \<longleftrightarrow> (\<forall>a\<in>A. b \<sqsubseteq> a)"
    1.14 -  by (auto intro: Inf_greatest dest: Inf_lower)
    1.15 +lemma Inf_mono:
    1.16 +  assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b"
    1.17 +  shows "Inf A \<le> Inf B"
    1.18 +proof (rule Inf_greatest)
    1.19 +  fix b assume "b \<in> B"
    1.20 +  with assms obtain a where "a \<in> A" and "a \<le> b" by blast
    1.21 +  from `a \<in> A` have "Inf A \<le> a" by (rule Inf_lower)
    1.22 +  with `a \<le> b` show "Inf A \<le> b" by auto
    1.23 +qed
    1.24  
    1.25  lemma Sup_mono:
    1.26    assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<le> b"
    1.27 @@ -98,49 +108,39 @@
    1.28    with `a \<le> b` show "a \<le> Sup B" by auto
    1.29  qed
    1.30  
    1.31 -lemma Inf_mono:
    1.32 -  assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b"
    1.33 -  shows "Inf A \<le> Inf B"
    1.34 -proof (rule Inf_greatest)
    1.35 -  fix b assume "b \<in> B"
    1.36 -  with assms obtain a where "a \<in> A" and "a \<le> b" by blast
    1.37 -  from `a \<in> A` have "Inf A \<le> a" by (rule Inf_lower)
    1.38 -  with `a \<le> b` show "Inf A \<le> b" by auto
    1.39 -qed
    1.40 +definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
    1.41 +  "INFI A f = \<Sqinter> (f ` A)"
    1.42  
    1.43  definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
    1.44    "SUPR A f = \<Squnion> (f ` A)"
    1.45  
    1.46 -definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
    1.47 -  "INFI A f = \<Sqinter> (f ` A)"
    1.48 -
    1.49  end
    1.50  
    1.51  syntax
    1.52 +  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
    1.53 +  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
    1.54    "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
    1.55    "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
    1.56 -  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
    1.57 -  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
    1.58  
    1.59  syntax (xsymbols)
    1.60 +  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
    1.61 +  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
    1.62    "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
    1.63    "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
    1.64 -  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
    1.65 -  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
    1.66  
    1.67  translations
    1.68 +  "INF x y. B"   == "INF x. INF y. B"
    1.69 +  "INF x. B"     == "CONST INFI CONST UNIV (%x. B)"
    1.70 +  "INF x. B"     == "INF x:CONST UNIV. B"
    1.71 +  "INF x:A. B"   == "CONST INFI A (%x. B)"
    1.72    "SUP x y. B"   == "SUP x. SUP y. B"
    1.73    "SUP x. B"     == "CONST SUPR CONST UNIV (%x. B)"
    1.74    "SUP x. B"     == "SUP x:CONST UNIV. B"
    1.75    "SUP x:A. B"   == "CONST SUPR A (%x. B)"
    1.76 -  "INF x y. B"   == "INF x. INF y. B"
    1.77 -  "INF x. B"     == "CONST INFI CONST UNIV (%x. B)"
    1.78 -  "INF x. B"     == "INF x:CONST UNIV. B"
    1.79 -  "INF x:A. B"   == "CONST INFI A (%x. B)"
    1.80  
    1.81  print_translation {*
    1.82 -  [Syntax.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"},
    1.83 -    Syntax.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"}]
    1.84 +  [Syntax.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"},
    1.85 +    Syntax.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}]
    1.86  *} -- {* to avoid eta-contraction of body *}
    1.87  
    1.88  context complete_lattice
    1.89 @@ -164,54 +164,54 @@
    1.90  lemma le_INF_iff: "u \<sqsubseteq> (INF i:A. M i) \<longleftrightarrow> (\<forall>i \<in> A. u \<sqsubseteq> M i)"
    1.91    unfolding INFI_def by (auto simp add: le_Inf_iff)
    1.92  
    1.93 +lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M"
    1.94 +  by (auto intro: antisym INF_leI le_INFI)
    1.95 +
    1.96  lemma SUP_const[simp]: "A \<noteq> {} \<Longrightarrow> (SUP i:A. M) = M"
    1.97    by (auto intro: antisym SUP_leI le_SUPI)
    1.98  
    1.99 -lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M"
   1.100 -  by (auto intro: antisym INF_leI le_INFI)
   1.101 +lemma INF_mono:
   1.102 +  "(\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> (INF n:A. f n) \<le> (INF n:B. g n)"
   1.103 +  by (force intro!: Inf_mono simp: INFI_def)
   1.104  
   1.105  lemma SUP_mono:
   1.106    "(\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> (SUP n:A. f n) \<le> (SUP n:B. g n)"
   1.107    by (force intro!: Sup_mono simp: SUPR_def)
   1.108  
   1.109 -lemma INF_mono:
   1.110 -  "(\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> (INF n:A. f n) \<le> (INF n:B. g n)"
   1.111 -  by (force intro!: Inf_mono simp: INFI_def)
   1.112 +lemma INF_subset:  "A \<subseteq> B \<Longrightarrow> INFI B f \<le> INFI A f"
   1.113 +  by (intro INF_mono) auto
   1.114  
   1.115  lemma SUP_subset:  "A \<subseteq> B \<Longrightarrow> SUPR A f \<le> SUPR B f"
   1.116    by (intro SUP_mono) auto
   1.117  
   1.118 -lemma INF_subset:  "A \<subseteq> B \<Longrightarrow> INFI B f \<le> INFI A f"
   1.119 -  by (intro INF_mono) auto
   1.120 +lemma INF_commute: "(INF i:A. INF j:B. f i j) = (INF j:B. INF i:A. f i j)"
   1.121 +  by (iprover intro: INF_leI le_INFI order_trans antisym)
   1.122  
   1.123  lemma SUP_commute: "(SUP i:A. SUP j:B. f i j) = (SUP j:B. SUP i:A. f i j)"
   1.124    by (iprover intro: SUP_leI le_SUPI order_trans antisym)
   1.125  
   1.126 -lemma INF_commute: "(INF i:A. INF j:B. f i j) = (INF j:B. INF i:A. f i j)"
   1.127 -  by (iprover intro: INF_leI le_INFI order_trans antisym)
   1.128 +end
   1.129  
   1.130 -end
   1.131 +lemma Inf_less_iff:
   1.132 +  fixes a :: "'a\<Colon>{complete_lattice,linorder}"
   1.133 +  shows "Inf S < a \<longleftrightarrow> (\<exists>x\<in>S. x < a)"
   1.134 +  unfolding not_le[symmetric] le_Inf_iff by auto
   1.135  
   1.136  lemma less_Sup_iff:
   1.137    fixes a :: "'a\<Colon>{complete_lattice,linorder}"
   1.138    shows "a < Sup S \<longleftrightarrow> (\<exists>x\<in>S. a < x)"
   1.139    unfolding not_le[symmetric] Sup_le_iff by auto
   1.140  
   1.141 -lemma Inf_less_iff:
   1.142 -  fixes a :: "'a\<Colon>{complete_lattice,linorder}"
   1.143 -  shows "Inf S < a \<longleftrightarrow> (\<exists>x\<in>S. x < a)"
   1.144 -  unfolding not_le[symmetric] le_Inf_iff by auto
   1.145 +lemma INF_less_iff:
   1.146 +  fixes a :: "'a::{complete_lattice,linorder}"
   1.147 +  shows "(INF i:A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
   1.148 +  unfolding INFI_def Inf_less_iff by auto
   1.149  
   1.150  lemma less_SUP_iff:
   1.151    fixes a :: "'a::{complete_lattice,linorder}"
   1.152    shows "a < (SUP i:A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
   1.153    unfolding SUPR_def less_Sup_iff by auto
   1.154  
   1.155 -lemma INF_less_iff:
   1.156 -  fixes a :: "'a::{complete_lattice,linorder}"
   1.157 -  shows "(INF i:A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
   1.158 -  unfolding INFI_def Inf_less_iff by auto
   1.159 -
   1.160  subsection {* @{typ bool} and @{typ "_ \<Rightarrow> _"} as complete lattice *}
   1.161  
   1.162  instantiation bool :: complete_lattice
   1.163 @@ -278,6 +278,200 @@
   1.164    by (auto intro: arg_cong [of _ _ Sup] simp add: SUPR_def Sup_apply)
   1.165  
   1.166  
   1.167 +subsection {* Inter *}
   1.168 +
   1.169 +abbreviation Inter :: "'a set set \<Rightarrow> 'a set" where
   1.170 +  "Inter S \<equiv> \<Sqinter>S"
   1.171 +  
   1.172 +notation (xsymbols)
   1.173 +  Inter  ("\<Inter>_" [90] 90)
   1.174 +
   1.175 +lemma Inter_eq:
   1.176 +  "\<Inter>A = {x. \<forall>B \<in> A. x \<in> B}"
   1.177 +proof (rule set_eqI)
   1.178 +  fix x
   1.179 +  have "(\<forall>Q\<in>{P. \<exists>B\<in>A. P \<longleftrightarrow> x \<in> B}. Q) \<longleftrightarrow> (\<forall>B\<in>A. x \<in> B)"
   1.180 +    by auto
   1.181 +  then show "x \<in> \<Inter>A \<longleftrightarrow> x \<in> {x. \<forall>B \<in> A. x \<in> B}"
   1.182 +    by (simp add: Inf_fun_def Inf_bool_def) (simp add: mem_def)
   1.183 +qed
   1.184 +
   1.185 +lemma Inter_iff [simp,no_atp]: "(A : Inter C) = (ALL X:C. A:X)"
   1.186 +  by (unfold Inter_eq) blast
   1.187 +
   1.188 +lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C"
   1.189 +  by (simp add: Inter_eq)
   1.190 +
   1.191 +text {*
   1.192 +  \medskip A ``destruct'' rule -- every @{term X} in @{term C}
   1.193 +  contains @{term A} as an element, but @{prop "A:X"} can hold when
   1.194 +  @{prop "X:C"} does not!  This rule is analogous to @{text spec}.
   1.195 +*}
   1.196 +
   1.197 +lemma InterD [elim, Pure.elim]: "A : Inter C ==> X:C ==> A:X"
   1.198 +  by auto
   1.199 +
   1.200 +lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R"
   1.201 +  -- {* ``Classical'' elimination rule -- does not require proving
   1.202 +    @{prop "X:C"}. *}
   1.203 +  by (unfold Inter_eq) blast
   1.204 +
   1.205 +lemma Inter_lower: "B \<in> A ==> Inter A \<subseteq> B"
   1.206 +  by blast
   1.207 +
   1.208 +lemma Inter_subset:
   1.209 +  "[| !!X. X \<in> A ==> X \<subseteq> B; A ~= {} |] ==> \<Inter>A \<subseteq> B"
   1.210 +  by blast
   1.211 +
   1.212 +lemma Inter_greatest: "(!!X. X \<in> A ==> C \<subseteq> X) ==> C \<subseteq> Inter A"
   1.213 +  by (iprover intro: InterI subsetI dest: subsetD)
   1.214 +
   1.215 +lemma Int_eq_Inter: "A \<inter> B = \<Inter>{A, B}"
   1.216 +  by blast
   1.217 +
   1.218 +lemma Inter_empty [simp]: "\<Inter>{} = UNIV"
   1.219 +  by (fact Inf_empty)
   1.220 +
   1.221 +lemma Inter_UNIV [simp]: "\<Inter>UNIV = {}"
   1.222 +  by blast
   1.223 +
   1.224 +lemma Inter_insert [simp]: "\<Inter>(insert a B) = a \<inter> \<Inter>B"
   1.225 +  by blast
   1.226 +
   1.227 +lemma Inter_Un_subset: "\<Inter>A \<union> \<Inter>B \<subseteq> \<Inter>(A \<inter> B)"
   1.228 +  by blast
   1.229 +
   1.230 +lemma Inter_Un_distrib: "\<Inter>(A \<union> B) = \<Inter>A \<inter> \<Inter>B"
   1.231 +  by blast
   1.232 +
   1.233 +lemma Inter_UNIV_conv [simp,no_atp]:
   1.234 +  "(\<Inter>A = UNIV) = (\<forall>x\<in>A. x = UNIV)"
   1.235 +  "(UNIV = \<Inter>A) = (\<forall>x\<in>A. x = UNIV)"
   1.236 +  by blast+
   1.237 +
   1.238 +lemma Inter_anti_mono: "B \<subseteq> A ==> \<Inter>A \<subseteq> \<Inter>B"
   1.239 +  by blast
   1.240 +
   1.241 +
   1.242 +subsection {* Intersections of families *}
   1.243 +
   1.244 +abbreviation INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
   1.245 +  "INTER \<equiv> INFI"
   1.246 +
   1.247 +syntax
   1.248 +  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3INT _./ _)" [0, 10] 10)
   1.249 +  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3INT _:_./ _)" [0, 0, 10] 10)
   1.250 +
   1.251 +syntax (xsymbols)
   1.252 +  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>_./ _)" [0, 10] 10)
   1.253 +  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>_\<in>_./ _)" [0, 0, 10] 10)
   1.254 +
   1.255 +syntax (latex output)
   1.256 +  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
   1.257 +  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
   1.258 +
   1.259 +translations
   1.260 +  "INT x y. B"  == "INT x. INT y. B"
   1.261 +  "INT x. B"    == "CONST INTER CONST UNIV (%x. B)"
   1.262 +  "INT x. B"    == "INT x:CONST UNIV. B"
   1.263 +  "INT x:A. B"  == "CONST INTER A (%x. B)"
   1.264 +
   1.265 +print_translation {*
   1.266 +  [Syntax.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}]
   1.267 +*} -- {* to avoid eta-contraction of body *}
   1.268 +
   1.269 +lemma INTER_eq_Inter_image:
   1.270 +  "(\<Inter>x\<in>A. B x) = \<Inter>(B`A)"
   1.271 +  by (fact INFI_def)
   1.272 +  
   1.273 +lemma Inter_def:
   1.274 +  "\<Inter>S = (\<Inter>x\<in>S. x)"
   1.275 +  by (simp add: INTER_eq_Inter_image image_def)
   1.276 +
   1.277 +lemma INTER_def:
   1.278 +  "(\<Inter>x\<in>A. B x) = {y. \<forall>x\<in>A. y \<in> B x}"
   1.279 +  by (auto simp add: INTER_eq_Inter_image Inter_eq)
   1.280 +
   1.281 +lemma Inter_image_eq [simp]:
   1.282 +  "\<Inter>(B`A) = (\<Inter>x\<in>A. B x)"
   1.283 +  by (rule sym) (fact INTER_eq_Inter_image)
   1.284 +
   1.285 +lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)"
   1.286 +  by (unfold INTER_def) blast
   1.287 +
   1.288 +lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)"
   1.289 +  by (unfold INTER_def) blast
   1.290 +
   1.291 +lemma INT_D [elim, Pure.elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a"
   1.292 +  by auto
   1.293 +
   1.294 +lemma INT_E [elim]: "b : (INT x:A. B x) ==> (b: B a ==> R) ==> (a~:A ==> R) ==> R"
   1.295 +  -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a:A"}. *}
   1.296 +  by (unfold INTER_def) blast
   1.297 +
   1.298 +lemma INT_cong [cong]:
   1.299 +    "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)"
   1.300 +  by (simp add: INTER_def)
   1.301 +
   1.302 +lemma Collect_ball_eq: "{x. \<forall>y\<in>A. P x y} = (\<Inter>y\<in>A. {x. P x y})"
   1.303 +  by blast
   1.304 +
   1.305 +lemma Collect_all_eq: "{x. \<forall>y. P x y} = (\<Inter>y. {x. P x y})"
   1.306 +  by blast
   1.307 +
   1.308 +lemma INT_lower: "a \<in> A ==> (\<Inter>x\<in>A. B x) \<subseteq> B a"
   1.309 +  by (fact INF_leI)
   1.310 +
   1.311 +lemma INT_greatest: "(!!x. x \<in> A ==> C \<subseteq> B x) ==> C \<subseteq> (\<Inter>x\<in>A. B x)"
   1.312 +  by (fact le_INFI)
   1.313 +
   1.314 +lemma INT_empty [simp]: "(\<Inter>x\<in>{}. B x) = UNIV"
   1.315 +  by blast
   1.316 +
   1.317 +lemma INT_absorb: "k \<in> I ==> A k \<inter> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. A i)"
   1.318 +  by blast
   1.319 +
   1.320 +lemma INT_subset_iff: "(B \<subseteq> (\<Inter>i\<in>I. A i)) = (\<forall>i\<in>I. B \<subseteq> A i)"
   1.321 +  by (fact le_INF_iff)
   1.322 +
   1.323 +lemma INT_insert [simp]: "(\<Inter>x \<in> insert a A. B x) = B a \<inter> INTER A B"
   1.324 +  by blast
   1.325 +
   1.326 +lemma INT_Un: "(\<Inter>i \<in> A \<union> B. M i) = (\<Inter>i \<in> A. M i) \<inter> (\<Inter>i\<in>B. M i)"
   1.327 +  by blast
   1.328 +
   1.329 +lemma INT_insert_distrib:
   1.330 +    "u \<in> A ==> (\<Inter>x\<in>A. insert a (B x)) = insert a (\<Inter>x\<in>A. B x)"
   1.331 +  by blast
   1.332 +
   1.333 +lemma INT_constant [simp]: "(\<Inter>y\<in>A. c) = (if A = {} then UNIV else c)"
   1.334 +  by auto
   1.335 +
   1.336 +lemma INT_eq: "(\<Inter>x\<in>A. B x) = \<Inter>({Y. \<exists>x\<in>A. Y = B x})"
   1.337 +  -- {* Look: it has an \emph{existential} quantifier *}
   1.338 +  by blast
   1.339 +
   1.340 +lemma INTER_UNIV_conv[simp]:
   1.341 + "(UNIV = (INT x:A. B x)) = (\<forall>x\<in>A. B x = UNIV)"
   1.342 + "((INT x:A. B x) = UNIV) = (\<forall>x\<in>A. B x = UNIV)"
   1.343 +by blast+
   1.344 +
   1.345 +lemma INT_bool_eq: "(\<Inter>b::bool. A b) = (A True \<inter> A False)"
   1.346 +  by (auto intro: bool_induct)
   1.347 +
   1.348 +lemma Pow_INT_eq: "Pow (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. Pow (B x))"
   1.349 +  by blast
   1.350 +
   1.351 +lemma INT_anti_mono:
   1.352 +  "B \<subseteq> A ==> (!!x. x \<in> A ==> f x \<subseteq> g x) ==>
   1.353 +    (\<Inter>x\<in>A. f x) \<subseteq> (\<Inter>x\<in>A. g x)"
   1.354 +  -- {* The last inclusion is POSITIVE! *}
   1.355 +  by (blast dest: subsetD)
   1.356 +
   1.357 +lemma vimage_INT: "f-`(INT x:A. B x) = (INT x:A. f -` B x)"
   1.358 +  by blast
   1.359 +
   1.360 +
   1.361  subsection {* Union *}
   1.362  
   1.363  abbreviation Union :: "'a set set \<Rightarrow> 'a set" where
   1.364 @@ -514,200 +708,6 @@
   1.365  by blast
   1.366  
   1.367  
   1.368 -subsection {* Inter *}
   1.369 -
   1.370 -abbreviation Inter :: "'a set set \<Rightarrow> 'a set" where
   1.371 -  "Inter S \<equiv> \<Sqinter>S"
   1.372 -  
   1.373 -notation (xsymbols)
   1.374 -  Inter  ("\<Inter>_" [90] 90)
   1.375 -
   1.376 -lemma Inter_eq:
   1.377 -  "\<Inter>A = {x. \<forall>B \<in> A. x \<in> B}"
   1.378 -proof (rule set_eqI)
   1.379 -  fix x
   1.380 -  have "(\<forall>Q\<in>{P. \<exists>B\<in>A. P \<longleftrightarrow> x \<in> B}. Q) \<longleftrightarrow> (\<forall>B\<in>A. x \<in> B)"
   1.381 -    by auto
   1.382 -  then show "x \<in> \<Inter>A \<longleftrightarrow> x \<in> {x. \<forall>B \<in> A. x \<in> B}"
   1.383 -    by (simp add: Inf_fun_def Inf_bool_def) (simp add: mem_def)
   1.384 -qed
   1.385 -
   1.386 -lemma Inter_iff [simp,no_atp]: "(A : Inter C) = (ALL X:C. A:X)"
   1.387 -  by (unfold Inter_eq) blast
   1.388 -
   1.389 -lemma InterI [intro!]: "(!!X. X:C ==> A:X) ==> A : Inter C"
   1.390 -  by (simp add: Inter_eq)
   1.391 -
   1.392 -text {*
   1.393 -  \medskip A ``destruct'' rule -- every @{term X} in @{term C}
   1.394 -  contains @{term A} as an element, but @{prop "A:X"} can hold when
   1.395 -  @{prop "X:C"} does not!  This rule is analogous to @{text spec}.
   1.396 -*}
   1.397 -
   1.398 -lemma InterD [elim, Pure.elim]: "A : Inter C ==> X:C ==> A:X"
   1.399 -  by auto
   1.400 -
   1.401 -lemma InterE [elim]: "A : Inter C ==> (X~:C ==> R) ==> (A:X ==> R) ==> R"
   1.402 -  -- {* ``Classical'' elimination rule -- does not require proving
   1.403 -    @{prop "X:C"}. *}
   1.404 -  by (unfold Inter_eq) blast
   1.405 -
   1.406 -lemma Inter_lower: "B \<in> A ==> Inter A \<subseteq> B"
   1.407 -  by blast
   1.408 -
   1.409 -lemma Inter_subset:
   1.410 -  "[| !!X. X \<in> A ==> X \<subseteq> B; A ~= {} |] ==> \<Inter>A \<subseteq> B"
   1.411 -  by blast
   1.412 -
   1.413 -lemma Inter_greatest: "(!!X. X \<in> A ==> C \<subseteq> X) ==> C \<subseteq> Inter A"
   1.414 -  by (iprover intro: InterI subsetI dest: subsetD)
   1.415 -
   1.416 -lemma Int_eq_Inter: "A \<inter> B = \<Inter>{A, B}"
   1.417 -  by blast
   1.418 -
   1.419 -lemma Inter_empty [simp]: "\<Inter>{} = UNIV"
   1.420 -  by (fact Inf_empty)
   1.421 -
   1.422 -lemma Inter_UNIV [simp]: "\<Inter>UNIV = {}"
   1.423 -  by blast
   1.424 -
   1.425 -lemma Inter_insert [simp]: "\<Inter>(insert a B) = a \<inter> \<Inter>B"
   1.426 -  by blast
   1.427 -
   1.428 -lemma Inter_Un_subset: "\<Inter>A \<union> \<Inter>B \<subseteq> \<Inter>(A \<inter> B)"
   1.429 -  by blast
   1.430 -
   1.431 -lemma Inter_Un_distrib: "\<Inter>(A \<union> B) = \<Inter>A \<inter> \<Inter>B"
   1.432 -  by blast
   1.433 -
   1.434 -lemma Inter_UNIV_conv [simp,no_atp]:
   1.435 -  "(\<Inter>A = UNIV) = (\<forall>x\<in>A. x = UNIV)"
   1.436 -  "(UNIV = \<Inter>A) = (\<forall>x\<in>A. x = UNIV)"
   1.437 -  by blast+
   1.438 -
   1.439 -lemma Inter_anti_mono: "B \<subseteq> A ==> \<Inter>A \<subseteq> \<Inter>B"
   1.440 -  by blast
   1.441 -
   1.442 -
   1.443 -subsection {* Intersections of families *}
   1.444 -
   1.445 -abbreviation INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
   1.446 -  "INTER \<equiv> INFI"
   1.447 -
   1.448 -syntax
   1.449 -  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3INT _./ _)" [0, 10] 10)
   1.450 -  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3INT _:_./ _)" [0, 0, 10] 10)
   1.451 -
   1.452 -syntax (xsymbols)
   1.453 -  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>_./ _)" [0, 10] 10)
   1.454 -  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>_\<in>_./ _)" [0, 0, 10] 10)
   1.455 -
   1.456 -syntax (latex output)
   1.457 -  "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
   1.458 -  "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
   1.459 -
   1.460 -translations
   1.461 -  "INT x y. B"  == "INT x. INT y. B"
   1.462 -  "INT x. B"    == "CONST INTER CONST UNIV (%x. B)"
   1.463 -  "INT x. B"    == "INT x:CONST UNIV. B"
   1.464 -  "INT x:A. B"  == "CONST INTER A (%x. B)"
   1.465 -
   1.466 -print_translation {*
   1.467 -  [Syntax.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}]
   1.468 -*} -- {* to avoid eta-contraction of body *}
   1.469 -
   1.470 -lemma INTER_eq_Inter_image:
   1.471 -  "(\<Inter>x\<in>A. B x) = \<Inter>(B`A)"
   1.472 -  by (fact INFI_def)
   1.473 -  
   1.474 -lemma Inter_def:
   1.475 -  "\<Inter>S = (\<Inter>x\<in>S. x)"
   1.476 -  by (simp add: INTER_eq_Inter_image image_def)
   1.477 -
   1.478 -lemma INTER_def:
   1.479 -  "(\<Inter>x\<in>A. B x) = {y. \<forall>x\<in>A. y \<in> B x}"
   1.480 -  by (auto simp add: INTER_eq_Inter_image Inter_eq)
   1.481 -
   1.482 -lemma Inter_image_eq [simp]:
   1.483 -  "\<Inter>(B`A) = (\<Inter>x\<in>A. B x)"
   1.484 -  by (rule sym) (fact INTER_eq_Inter_image)
   1.485 -
   1.486 -lemma INT_iff [simp]: "(b: (INT x:A. B x)) = (ALL x:A. b: B x)"
   1.487 -  by (unfold INTER_def) blast
   1.488 -
   1.489 -lemma INT_I [intro!]: "(!!x. x:A ==> b: B x) ==> b : (INT x:A. B x)"
   1.490 -  by (unfold INTER_def) blast
   1.491 -
   1.492 -lemma INT_D [elim, Pure.elim]: "b : (INT x:A. B x) ==> a:A ==> b: B a"
   1.493 -  by auto
   1.494 -
   1.495 -lemma INT_E [elim]: "b : (INT x:A. B x) ==> (b: B a ==> R) ==> (a~:A ==> R) ==> R"
   1.496 -  -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a:A"}. *}
   1.497 -  by (unfold INTER_def) blast
   1.498 -
   1.499 -lemma INT_cong [cong]:
   1.500 -    "A = B ==> (!!x. x:B ==> C x = D x) ==> (INT x:A. C x) = (INT x:B. D x)"
   1.501 -  by (simp add: INTER_def)
   1.502 -
   1.503 -lemma Collect_ball_eq: "{x. \<forall>y\<in>A. P x y} = (\<Inter>y\<in>A. {x. P x y})"
   1.504 -  by blast
   1.505 -
   1.506 -lemma Collect_all_eq: "{x. \<forall>y. P x y} = (\<Inter>y. {x. P x y})"
   1.507 -  by blast
   1.508 -
   1.509 -lemma INT_lower: "a \<in> A ==> (\<Inter>x\<in>A. B x) \<subseteq> B a"
   1.510 -  by (fact INF_leI)
   1.511 -
   1.512 -lemma INT_greatest: "(!!x. x \<in> A ==> C \<subseteq> B x) ==> C \<subseteq> (\<Inter>x\<in>A. B x)"
   1.513 -  by (fact le_INFI)
   1.514 -
   1.515 -lemma INT_empty [simp]: "(\<Inter>x\<in>{}. B x) = UNIV"
   1.516 -  by blast
   1.517 -
   1.518 -lemma INT_absorb: "k \<in> I ==> A k \<inter> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. A i)"
   1.519 -  by blast
   1.520 -
   1.521 -lemma INT_subset_iff: "(B \<subseteq> (\<Inter>i\<in>I. A i)) = (\<forall>i\<in>I. B \<subseteq> A i)"
   1.522 -  by (fact le_INF_iff)
   1.523 -
   1.524 -lemma INT_insert [simp]: "(\<Inter>x \<in> insert a A. B x) = B a \<inter> INTER A B"
   1.525 -  by blast
   1.526 -
   1.527 -lemma INT_Un: "(\<Inter>i \<in> A \<union> B. M i) = (\<Inter>i \<in> A. M i) \<inter> (\<Inter>i\<in>B. M i)"
   1.528 -  by blast
   1.529 -
   1.530 -lemma INT_insert_distrib:
   1.531 -    "u \<in> A ==> (\<Inter>x\<in>A. insert a (B x)) = insert a (\<Inter>x\<in>A. B x)"
   1.532 -  by blast
   1.533 -
   1.534 -lemma INT_constant [simp]: "(\<Inter>y\<in>A. c) = (if A = {} then UNIV else c)"
   1.535 -  by auto
   1.536 -
   1.537 -lemma INT_eq: "(\<Inter>x\<in>A. B x) = \<Inter>({Y. \<exists>x\<in>A. Y = B x})"
   1.538 -  -- {* Look: it has an \emph{existential} quantifier *}
   1.539 -  by blast
   1.540 -
   1.541 -lemma INTER_UNIV_conv[simp]:
   1.542 - "(UNIV = (INT x:A. B x)) = (\<forall>x\<in>A. B x = UNIV)"
   1.543 - "((INT x:A. B x) = UNIV) = (\<forall>x\<in>A. B x = UNIV)"
   1.544 -by blast+
   1.545 -
   1.546 -lemma INT_bool_eq: "(\<Inter>b::bool. A b) = (A True \<inter> A False)"
   1.547 -  by (auto intro: bool_induct)
   1.548 -
   1.549 -lemma Pow_INT_eq: "Pow (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. Pow (B x))"
   1.550 -  by blast
   1.551 -
   1.552 -lemma INT_anti_mono:
   1.553 -  "B \<subseteq> A ==> (!!x. x \<in> A ==> f x \<subseteq> g x) ==>
   1.554 -    (\<Inter>x\<in>A. f x) \<subseteq> (\<Inter>x\<in>A. g x)"
   1.555 -  -- {* The last inclusion is POSITIVE! *}
   1.556 -  by (blast dest: subsetD)
   1.557 -
   1.558 -lemma vimage_INT: "f-`(INT x:A. B x) = (INT x:A. f -` B x)"
   1.559 -  by blast
   1.560 -
   1.561 -
   1.562  subsection {* Distributive laws *}
   1.563  
   1.564  lemma Int_Union: "A \<inter> \<Union>B = (\<Union>C\<in>B. A \<inter> C)"
   1.565 @@ -858,18 +858,18 @@
   1.566  no_notation
   1.567    less_eq  (infix "\<sqsubseteq>" 50) and
   1.568    less (infix "\<sqsubset>" 50) and
   1.569 +  bot ("\<bottom>") and
   1.570 +  top ("\<top>") and
   1.571    inf  (infixl "\<sqinter>" 70) and
   1.572    sup  (infixl "\<squnion>" 65) and
   1.573    Inf  ("\<Sqinter>_" [900] 900) and
   1.574 -  Sup  ("\<Squnion>_" [900] 900) and
   1.575 -  top ("\<top>") and
   1.576 -  bot ("\<bottom>")
   1.577 +  Sup  ("\<Squnion>_" [900] 900)
   1.578  
   1.579  no_syntax (xsymbols)
   1.580 +  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
   1.581 +  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
   1.582    "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
   1.583    "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
   1.584 -  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
   1.585 -  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
   1.586  
   1.587  lemmas mem_simps =
   1.588    insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff
     2.1 --- a/src/HOL/Lattices.thy	Wed Dec 08 14:52:23 2010 +0100
     2.2 +++ b/src/HOL/Lattices.thy	Wed Dec 08 15:05:46 2010 +0100
     2.3 @@ -48,8 +48,9 @@
     2.4  notation
     2.5    less_eq  (infix "\<sqsubseteq>" 50) and
     2.6    less  (infix "\<sqsubset>" 50) and
     2.7 -  top ("\<top>") and
     2.8 -  bot ("\<bottom>")
     2.9 +  bot ("\<bottom>") and
    2.10 +  top ("\<top>")
    2.11 +
    2.12  
    2.13  class semilattice_inf = order +
    2.14    fixes inf :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70)
     3.1 --- a/src/HOL/Library/Lattice_Syntax.thy	Wed Dec 08 14:52:23 2010 +0100
     3.2 +++ b/src/HOL/Library/Lattice_Syntax.thy	Wed Dec 08 15:05:46 2010 +0100
     3.3 @@ -7,17 +7,17 @@
     3.4  begin
     3.5  
     3.6  notation
     3.7 +  bot ("\<bottom>") and
     3.8    top ("\<top>") and
     3.9 -  bot ("\<bottom>") and
    3.10    inf  (infixl "\<sqinter>" 70) and
    3.11    sup  (infixl "\<squnion>" 65) and
    3.12    Inf  ("\<Sqinter>_" [900] 900) and
    3.13    Sup  ("\<Squnion>_" [900] 900)
    3.14  
    3.15  syntax (xsymbols)
    3.16 +  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
    3.17 +  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
    3.18    "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
    3.19    "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
    3.20 -  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
    3.21 -  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
    3.22  
    3.23  end
     4.1 --- a/src/HOL/Orderings.thy	Wed Dec 08 14:52:23 2010 +0100
     4.2 +++ b/src/HOL/Orderings.thy	Wed Dec 08 15:05:46 2010 +0100
     4.3 @@ -1082,14 +1082,14 @@
     4.4  
     4.5  subsection {* Top and bottom elements *}
     4.6  
     4.7 +class bot = preorder +
     4.8 +  fixes bot :: 'a
     4.9 +  assumes bot_least [simp]: "bot \<le> x"
    4.10 +
    4.11  class top = preorder +
    4.12    fixes top :: 'a
    4.13    assumes top_greatest [simp]: "x \<le> top"
    4.14  
    4.15 -class bot = preorder +
    4.16 -  fixes bot :: 'a
    4.17 -  assumes bot_least [simp]: "bot \<le> x"
    4.18 -
    4.19  
    4.20  subsection {* Dense orders *}
    4.21  
    4.22 @@ -1204,7 +1204,7 @@
    4.23  
    4.24  subsection {* Order on bool *}
    4.25  
    4.26 -instantiation bool :: "{order, top, bot}"
    4.27 +instantiation bool :: "{order, bot, top}"
    4.28  begin
    4.29  
    4.30  definition
    4.31 @@ -1214,10 +1214,10 @@
    4.32    [simp]: "(P\<Colon>bool) < Q \<longleftrightarrow> \<not> P \<and> Q"
    4.33  
    4.34  definition
    4.35 -  [simp]: "top \<longleftrightarrow> True"
    4.36 +  [simp]: "bot \<longleftrightarrow> False"
    4.37  
    4.38  definition
    4.39 -  [simp]: "bot \<longleftrightarrow> False"
    4.40 +  [simp]: "top \<longleftrightarrow> True"
    4.41  
    4.42  instance proof
    4.43  qed auto
    4.44 @@ -1272,6 +1272,21 @@
    4.45  instance "fun" :: (type, order) order proof
    4.46  qed (auto simp add: le_fun_def intro: antisym ext)
    4.47  
    4.48 +instantiation "fun" :: (type, bot) bot
    4.49 +begin
    4.50 +
    4.51 +definition
    4.52 +  "bot = (\<lambda>x. bot)"
    4.53 +
    4.54 +lemma bot_apply:
    4.55 +  "bot x = bot"
    4.56 +  by (simp add: bot_fun_def)
    4.57 +
    4.58 +instance proof
    4.59 +qed (simp add: le_fun_def bot_apply)
    4.60 +
    4.61 +end
    4.62 +
    4.63  instantiation "fun" :: (type, top) top
    4.64  begin
    4.65  
    4.66 @@ -1288,21 +1303,6 @@
    4.67  
    4.68  end
    4.69  
    4.70 -instantiation "fun" :: (type, bot) bot
    4.71 -begin
    4.72 -
    4.73 -definition
    4.74 -  "bot = (\<lambda>x. bot)"
    4.75 -
    4.76 -lemma bot_apply:
    4.77 -  "bot x = bot"
    4.78 -  by (simp add: bot_fun_def)
    4.79 -
    4.80 -instance proof
    4.81 -qed (simp add: le_fun_def bot_apply)
    4.82 -
    4.83 -end
    4.84 -
    4.85  lemma le_funI: "(\<And>x. f x \<le> g x) \<Longrightarrow> f \<le> g"
    4.86    unfolding le_fun_def by simp
    4.87  
     5.1 --- a/src/HOL/Predicate.thy	Wed Dec 08 14:52:23 2010 +0100
     5.2 +++ b/src/HOL/Predicate.thy	Wed Dec 08 15:05:46 2010 +0100
     5.3 @@ -9,18 +9,18 @@
     5.4  begin
     5.5  
     5.6  notation
     5.7 +  bot ("\<bottom>") and
     5.8 +  top ("\<top>") and
     5.9    inf (infixl "\<sqinter>" 70) and
    5.10    sup (infixl "\<squnion>" 65) and
    5.11    Inf ("\<Sqinter>_" [900] 900) and
    5.12 -  Sup ("\<Squnion>_" [900] 900) and
    5.13 -  top ("\<top>") and
    5.14 -  bot ("\<bottom>")
    5.15 +  Sup ("\<Squnion>_" [900] 900)
    5.16  
    5.17  syntax (xsymbols)
    5.18 +  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
    5.19 +  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
    5.20    "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
    5.21    "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
    5.22 -  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
    5.23 -  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
    5.24  
    5.25  
    5.26  subsection {* Predicates as (complete) lattices *}
    5.27 @@ -92,12 +92,6 @@
    5.28  
    5.29  subsubsection {* Top and bottom elements *}
    5.30  
    5.31 -lemma top1I [intro!]: "top x"
    5.32 -  by (simp add: top_fun_def top_bool_def)
    5.33 -
    5.34 -lemma top2I [intro!]: "top x y"
    5.35 -  by (simp add: top_fun_def top_bool_def)
    5.36 -
    5.37  lemma bot1E [no_atp, elim!]: "bot x \<Longrightarrow> P"
    5.38    by (simp add: bot_fun_def bot_bool_def)
    5.39  
    5.40 @@ -110,6 +104,45 @@
    5.41  lemma bot_empty_eq2: "bot = (\<lambda>x y. (x, y) \<in> {})"
    5.42    by (auto simp add: fun_eq_iff)
    5.43  
    5.44 +lemma top1I [intro!]: "top x"
    5.45 +  by (simp add: top_fun_def top_bool_def)
    5.46 +
    5.47 +lemma top2I [intro!]: "top x y"
    5.48 +  by (simp add: top_fun_def top_bool_def)
    5.49 +
    5.50 +
    5.51 +subsubsection {* Binary intersection *}
    5.52 +
    5.53 +lemma inf1I [intro!]: "A x ==> B x ==> inf A B x"
    5.54 +  by (simp add: inf_fun_def inf_bool_def)
    5.55 +
    5.56 +lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y"
    5.57 +  by (simp add: inf_fun_def inf_bool_def)
    5.58 +
    5.59 +lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P"
    5.60 +  by (simp add: inf_fun_def inf_bool_def)
    5.61 +
    5.62 +lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P"
    5.63 +  by (simp add: inf_fun_def inf_bool_def)
    5.64 +
    5.65 +lemma inf1D1: "inf A B x ==> A x"
    5.66 +  by (simp add: inf_fun_def inf_bool_def)
    5.67 +
    5.68 +lemma inf2D1: "inf A B x y ==> A x y"
    5.69 +  by (simp add: inf_fun_def inf_bool_def)
    5.70 +
    5.71 +lemma inf1D2: "inf A B x ==> B x"
    5.72 +  by (simp add: inf_fun_def inf_bool_def)
    5.73 +
    5.74 +lemma inf2D2: "inf A B x y ==> B x y"
    5.75 +  by (simp add: inf_fun_def inf_bool_def)
    5.76 +
    5.77 +lemma inf_Int_eq: "inf (\<lambda>x. x \<in> R) (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<inter> S)"
    5.78 +  by (simp add: inf_fun_def inf_bool_def mem_def)
    5.79 +
    5.80 +lemma inf_Int_eq2 [pred_set_conv]: "inf (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<inter> S)"
    5.81 +  by (simp add: inf_fun_def inf_bool_def mem_def)
    5.82 +
    5.83  
    5.84  subsubsection {* Binary union *}
    5.85  
    5.86 @@ -149,66 +182,6 @@
    5.87    by (simp add: sup_fun_def sup_bool_def mem_def)
    5.88  
    5.89  
    5.90 -subsubsection {* Binary intersection *}
    5.91 -
    5.92 -lemma inf1I [intro!]: "A x ==> B x ==> inf A B x"
    5.93 -  by (simp add: inf_fun_def inf_bool_def)
    5.94 -
    5.95 -lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y"
    5.96 -  by (simp add: inf_fun_def inf_bool_def)
    5.97 -
    5.98 -lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P"
    5.99 -  by (simp add: inf_fun_def inf_bool_def)
   5.100 -
   5.101 -lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P"
   5.102 -  by (simp add: inf_fun_def inf_bool_def)
   5.103 -
   5.104 -lemma inf1D1: "inf A B x ==> A x"
   5.105 -  by (simp add: inf_fun_def inf_bool_def)
   5.106 -
   5.107 -lemma inf2D1: "inf A B x y ==> A x y"
   5.108 -  by (simp add: inf_fun_def inf_bool_def)
   5.109 -
   5.110 -lemma inf1D2: "inf A B x ==> B x"
   5.111 -  by (simp add: inf_fun_def inf_bool_def)
   5.112 -
   5.113 -lemma inf2D2: "inf A B x y ==> B x y"
   5.114 -  by (simp add: inf_fun_def inf_bool_def)
   5.115 -
   5.116 -lemma inf_Int_eq: "inf (\<lambda>x. x \<in> R) (\<lambda>x. x \<in> S) = (\<lambda>x. x \<in> R \<inter> S)"
   5.117 -  by (simp add: inf_fun_def inf_bool_def mem_def)
   5.118 -
   5.119 -lemma inf_Int_eq2 [pred_set_conv]: "inf (\<lambda>x y. (x, y) \<in> R) (\<lambda>x y. (x, y) \<in> S) = (\<lambda>x y. (x, y) \<in> R \<inter> S)"
   5.120 -  by (simp add: inf_fun_def inf_bool_def mem_def)
   5.121 -
   5.122 -
   5.123 -subsubsection {* Unions of families *}
   5.124 -
   5.125 -lemma SUP1_iff: "(SUP x:A. B x) b = (EX x:A. B x b)"
   5.126 -  by (simp add: SUPR_apply)
   5.127 -
   5.128 -lemma SUP2_iff: "(SUP x:A. B x) b c = (EX x:A. B x b c)"
   5.129 -  by (simp add: SUPR_apply)
   5.130 -
   5.131 -lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b"
   5.132 -  by (auto simp add: SUPR_apply)
   5.133 -
   5.134 -lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c"
   5.135 -  by (auto simp add: SUPR_apply)
   5.136 -
   5.137 -lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R"
   5.138 -  by (auto simp add: SUPR_apply)
   5.139 -
   5.140 -lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R"
   5.141 -  by (auto simp add: SUPR_apply)
   5.142 -
   5.143 -lemma SUP_UN_eq: "(SUP i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (UN i. r i))"
   5.144 -  by (simp add: SUPR_apply fun_eq_iff)
   5.145 -
   5.146 -lemma SUP_UN_eq2: "(SUP i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (UN i. r i))"
   5.147 -  by (simp add: SUPR_apply fun_eq_iff)
   5.148 -
   5.149 -
   5.150  subsubsection {* Intersections of families *}
   5.151  
   5.152  lemma INF1_iff: "(INF x:A. B x) b = (ALL x:A. B x b)"
   5.153 @@ -242,6 +215,33 @@
   5.154    by (simp add: INFI_apply fun_eq_iff)
   5.155  
   5.156  
   5.157 +subsubsection {* Unions of families *}
   5.158 +
   5.159 +lemma SUP1_iff: "(SUP x:A. B x) b = (EX x:A. B x b)"
   5.160 +  by (simp add: SUPR_apply)
   5.161 +
   5.162 +lemma SUP2_iff: "(SUP x:A. B x) b c = (EX x:A. B x b c)"
   5.163 +  by (simp add: SUPR_apply)
   5.164 +
   5.165 +lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b"
   5.166 +  by (auto simp add: SUPR_apply)
   5.167 +
   5.168 +lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c"
   5.169 +  by (auto simp add: SUPR_apply)
   5.170 +
   5.171 +lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R"
   5.172 +  by (auto simp add: SUPR_apply)
   5.173 +
   5.174 +lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R"
   5.175 +  by (auto simp add: SUPR_apply)
   5.176 +
   5.177 +lemma SUP_UN_eq: "(SUP i. (\<lambda>x. x \<in> r i)) = (\<lambda>x. x \<in> (UN i. r i))"
   5.178 +  by (simp add: SUPR_apply fun_eq_iff)
   5.179 +
   5.180 +lemma SUP_UN_eq2: "(SUP i. (\<lambda>x y. (x, y) \<in> r i)) = (\<lambda>x y. (x, y) \<in> (UN i. r i))"
   5.181 +  by (simp add: SUPR_apply fun_eq_iff)
   5.182 +
   5.183 +
   5.184  subsection {* Predicates as relations *}
   5.185  
   5.186  subsubsection {* Composition  *}
   5.187 @@ -1027,19 +1027,19 @@
   5.188  *}
   5.189  
   5.190  no_notation
   5.191 +  bot ("\<bottom>") and
   5.192 +  top ("\<top>") and
   5.193    inf (infixl "\<sqinter>" 70) and
   5.194    sup (infixl "\<squnion>" 65) and
   5.195    Inf ("\<Sqinter>_" [900] 900) and
   5.196    Sup ("\<Squnion>_" [900] 900) and
   5.197 -  top ("\<top>") and
   5.198 -  bot ("\<bottom>") and
   5.199    bind (infixl "\<guillemotright>=" 70)
   5.200  
   5.201  no_syntax (xsymbols)
   5.202 +  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
   5.203 +  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
   5.204    "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
   5.205    "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
   5.206 -  "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
   5.207 -  "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
   5.208  
   5.209  hide_type (open) pred seq
   5.210  hide_const (open) Pred eval single bind is_empty singleton if_pred not_pred holds
     6.1 --- a/src/HOL/Set.thy	Wed Dec 08 14:52:23 2010 +0100
     6.2 +++ b/src/HOL/Set.thy	Wed Dec 08 15:05:46 2010 +0100
     6.3 @@ -533,6 +533,36 @@
     6.4    by simp
     6.5  
     6.6  
     6.7 +subsubsection {* The empty set *}
     6.8 +
     6.9 +lemma empty_def:
    6.10 +  "{} = {x. False}"
    6.11 +  by (simp add: bot_fun_def bot_bool_def Collect_def)
    6.12 +
    6.13 +lemma empty_iff [simp]: "(c : {}) = False"
    6.14 +  by (simp add: empty_def)
    6.15 +
    6.16 +lemma emptyE [elim!]: "a : {} ==> P"
    6.17 +  by simp
    6.18 +
    6.19 +lemma empty_subsetI [iff]: "{} \<subseteq> A"
    6.20 +    -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *}
    6.21 +  by blast
    6.22 +
    6.23 +lemma equals0I: "(!!y. y \<in> A ==> False) ==> A = {}"
    6.24 +  by blast
    6.25 +
    6.26 +lemma equals0D: "A = {} ==> a \<notin> A"
    6.27 +    -- {* Use for reasoning about disjointness: @{text "A Int B = {}"} *}
    6.28 +  by blast
    6.29 +
    6.30 +lemma ball_empty [simp]: "Ball {} P = True"
    6.31 +  by (simp add: Ball_def)
    6.32 +
    6.33 +lemma bex_empty [simp]: "Bex {} P = False"
    6.34 +  by (simp add: Bex_def)
    6.35 +
    6.36 +
    6.37  subsubsection {* The universal set -- UNIV *}
    6.38  
    6.39  abbreviation UNIV :: "'a set" where
    6.40 @@ -568,36 +598,6 @@
    6.41  lemma UNIV_eq_I: "(\<And>x. x \<in> A) \<Longrightarrow> UNIV = A"
    6.42    by auto
    6.43  
    6.44 -
    6.45 -subsubsection {* The empty set *}
    6.46 -
    6.47 -lemma empty_def:
    6.48 -  "{} = {x. False}"
    6.49 -  by (simp add: bot_fun_def bot_bool_def Collect_def)
    6.50 -
    6.51 -lemma empty_iff [simp]: "(c : {}) = False"
    6.52 -  by (simp add: empty_def)
    6.53 -
    6.54 -lemma emptyE [elim!]: "a : {} ==> P"
    6.55 -  by simp
    6.56 -
    6.57 -lemma empty_subsetI [iff]: "{} \<subseteq> A"
    6.58 -    -- {* One effect is to delete the ASSUMPTION @{prop "{} <= A"} *}
    6.59 -  by blast
    6.60 -
    6.61 -lemma equals0I: "(!!y. y \<in> A ==> False) ==> A = {}"
    6.62 -  by blast
    6.63 -
    6.64 -lemma equals0D: "A = {} ==> a \<notin> A"
    6.65 -    -- {* Use for reasoning about disjointness: @{text "A Int B = {}"} *}
    6.66 -  by blast
    6.67 -
    6.68 -lemma ball_empty [simp]: "Ball {} P = True"
    6.69 -  by (simp add: Ball_def)
    6.70 -
    6.71 -lemma bex_empty [simp]: "Bex {} P = False"
    6.72 -  by (simp add: Bex_def)
    6.73 -
    6.74  lemma UNIV_not_empty [iff]: "UNIV ~= {}"
    6.75    by (blast elim: equalityE)
    6.76  
    6.77 @@ -647,7 +647,41 @@
    6.78  lemma Compl_eq: "- A = {x. ~ x : A}" by blast
    6.79  
    6.80  
    6.81 -subsubsection {* Binary union -- Un *}
    6.82 +subsubsection {* Binary intersection *}
    6.83 +
    6.84 +abbreviation inter :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Int" 70) where
    6.85 +  "op Int \<equiv> inf"
    6.86 +
    6.87 +notation (xsymbols)
    6.88 +  inter  (infixl "\<inter>" 70)
    6.89 +
    6.90 +notation (HTML output)
    6.91 +  inter  (infixl "\<inter>" 70)
    6.92 +
    6.93 +lemma Int_def:
    6.94 +  "A \<inter> B = {x. x \<in> A \<and> x \<in> B}"
    6.95 +  by (simp add: inf_fun_def inf_bool_def Collect_def mem_def)
    6.96 +
    6.97 +lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)"
    6.98 +  by (unfold Int_def) blast
    6.99 +
   6.100 +lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B"
   6.101 +  by simp
   6.102 +
   6.103 +lemma IntD1: "c : A Int B ==> c:A"
   6.104 +  by simp
   6.105 +
   6.106 +lemma IntD2: "c : A Int B ==> c:B"
   6.107 +  by simp
   6.108 +
   6.109 +lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P"
   6.110 +  by simp
   6.111 +
   6.112 +lemma mono_Int: "mono f \<Longrightarrow> f (A \<inter> B) \<subseteq> f A \<inter> f B"
   6.113 +  by (fact mono_inf)
   6.114 +
   6.115 +
   6.116 +subsubsection {* Binary union *}
   6.117  
   6.118  abbreviation union :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Un" 65) where
   6.119    "union \<equiv> sup"
   6.120 @@ -689,40 +723,6 @@
   6.121    by (fact mono_sup)
   6.122  
   6.123  
   6.124 -subsubsection {* Binary intersection -- Int *}
   6.125 -
   6.126 -abbreviation inter :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Int" 70) where
   6.127 -  "op Int \<equiv> inf"
   6.128 -
   6.129 -notation (xsymbols)
   6.130 -  inter  (infixl "\<inter>" 70)
   6.131 -
   6.132 -notation (HTML output)
   6.133 -  inter  (infixl "\<inter>" 70)
   6.134 -
   6.135 -lemma Int_def:
   6.136 -  "A \<inter> B = {x. x \<in> A \<and> x \<in> B}"
   6.137 -  by (simp add: inf_fun_def inf_bool_def Collect_def mem_def)
   6.138 -
   6.139 -lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)"
   6.140 -  by (unfold Int_def) blast
   6.141 -
   6.142 -lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B"
   6.143 -  by simp
   6.144 -
   6.145 -lemma IntD1: "c : A Int B ==> c:A"
   6.146 -  by simp
   6.147 -
   6.148 -lemma IntD2: "c : A Int B ==> c:B"
   6.149 -  by simp
   6.150 -
   6.151 -lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P"
   6.152 -  by simp
   6.153 -
   6.154 -lemma mono_Int: "mono f \<Longrightarrow> f (A \<inter> B) \<subseteq> f A \<inter> f B"
   6.155 -  by (fact mono_inf)
   6.156 -
   6.157 -
   6.158  subsubsection {* Set difference *}
   6.159  
   6.160  lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)"