src/HOL/Complete_Lattices.thy
author haftmann
Thu, 25 Jul 2013 08:57:16 +0200
changeset 53866 412c9e0381a1
parent 53278 eff000cab70f
child 54152 a1119cf551e8
permissions -rw-r--r--
factored syntactic type classes for bot and top (by Alessandro Coglio)
     1  (*  Author:     Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *)
     2 
     3 header {* Complete lattices *}
     4 
     5 theory Complete_Lattices
     6 imports Set
     7 begin
     8 
     9 notation
    10   less_eq (infix "\<sqsubseteq>" 50) and
    11   less (infix "\<sqsubset>" 50)
    12 
    13 
    14 subsection {* Syntactic infimum and supremum operations *}
    15 
    16 class Inf =
    17   fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
    18 
    19 class Sup =
    20   fixes Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
    21 
    22 
    23 subsection {* Abstract complete lattices *}
    24 
    25 text {* A complete lattice always has a bottom and a top,
    26 so we include them into the following type class,
    27 along with assumptions that define bottom and top
    28 in terms of infimum and supremum. *}
    29 
    30 class complete_lattice = lattice + Inf + Sup + bot + top +
    31   assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x"
    32      and Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A"
    33   assumes Sup_upper: "x \<in> A \<Longrightarrow> x \<sqsubseteq> \<Squnion>A"
    34      and Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> z) \<Longrightarrow> \<Squnion>A \<sqsubseteq> z"
    35   assumes Inf_empty [simp]: "\<Sqinter>{} = \<top>"
    36   assumes Sup_empty [simp]: "\<Squnion>{} = \<bottom>"
    37 begin
    38 
    39 subclass bounded_lattice
    40 proof
    41   fix a
    42   show "\<bottom> \<le> a" by (auto intro: Sup_least simp only: Sup_empty [symmetric])
    43   show "a \<le> \<top>" by (auto intro: Inf_greatest simp only: Inf_empty [symmetric])
    44 qed
    45 
    46 lemma dual_complete_lattice:
    47   "class.complete_lattice Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom>"
    48   by (auto intro!: class.complete_lattice.intro dual_lattice)
    49     (unfold_locales, (fact Inf_empty Sup_empty
    50         Sup_upper Sup_least Inf_lower Inf_greatest)+)
    51 
    52 definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
    53   INF_def: "INFI A f = \<Sqinter>(f ` A)"
    54 
    55 definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
    56   SUP_def: "SUPR A f = \<Squnion>(f ` A)"
    57 
    58 text {*
    59   Note: must use names @{const INFI} and @{const SUPR} here instead of
    60   @{text INF} and @{text SUP} to allow the following syntax coexist
    61   with the plain constant names.
    62 *}
    63 
    64 end
    65 
    66 syntax
    67   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3INF _./ _)" [0, 10] 10)
    68   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3INF _:_./ _)" [0, 0, 10] 10)
    69   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3SUP _./ _)" [0, 10] 10)
    70   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3SUP _:_./ _)" [0, 0, 10] 10)
    71 
    72 syntax (xsymbols)
    73   "_INF1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Sqinter>_./ _)" [0, 10] 10)
    74   "_INF"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
    75   "_SUP1"     :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b"           ("(3\<Squnion>_./ _)" [0, 10] 10)
    76   "_SUP"      :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"  ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
    77 
    78 translations
    79   "INF x y. B"   == "INF x. INF y. B"
    80   "INF x. B"     == "CONST INFI CONST UNIV (%x. B)"
    81   "INF x. B"     == "INF x:CONST UNIV. B"
    82   "INF x:A. B"   == "CONST INFI A (%x. B)"
    83   "SUP x y. B"   == "SUP x. SUP y. B"
    84   "SUP x. B"     == "CONST SUPR CONST UNIV (%x. B)"
    85   "SUP x. B"     == "SUP x:CONST UNIV. B"
    86   "SUP x:A. B"   == "CONST SUPR A (%x. B)"
    87 
    88 print_translation {*
    89   [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"},
    90     Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}]
    91 *} -- {* to avoid eta-contraction of body *}
    92 
    93 context complete_lattice
    94 begin
    95 
    96 lemma INF_foundation_dual [no_atp]:
    97   "complete_lattice.SUPR Inf = INFI"
    98   by (simp add: fun_eq_iff INF_def
    99     complete_lattice.SUP_def [OF dual_complete_lattice])
   100 
   101 lemma SUP_foundation_dual [no_atp]:
   102   "complete_lattice.INFI Sup = SUPR"
   103   by (simp add: fun_eq_iff SUP_def
   104     complete_lattice.INF_def [OF dual_complete_lattice])
   105 
   106 lemma Sup_eqI:
   107   "(\<And>y. y \<in> A \<Longrightarrow> y \<le> x) \<Longrightarrow> (\<And>y. (\<And>z. z \<in> A \<Longrightarrow> z \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> \<Squnion>A = x"
   108   by (blast intro: antisym Sup_least Sup_upper)
   109 
   110 lemma Inf_eqI:
   111   "(\<And>i. i \<in> A \<Longrightarrow> x \<le> i) \<Longrightarrow> (\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> i) \<Longrightarrow> y \<le> x) \<Longrightarrow> \<Sqinter>A = x"
   112   by (blast intro: antisym Inf_greatest Inf_lower)
   113 
   114 lemma SUP_eqI:
   115   "(\<And>i. i \<in> A \<Longrightarrow> f i \<le> x) \<Longrightarrow> (\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> x \<le> y) \<Longrightarrow> (\<Squnion>i\<in>A. f i) = x"
   116   unfolding SUP_def by (rule Sup_eqI) auto
   117 
   118 lemma INF_eqI:
   119   "(\<And>i. i \<in> A \<Longrightarrow> x \<le> f i) \<Longrightarrow> (\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<ge> y) \<Longrightarrow> x \<ge> y) \<Longrightarrow> (\<Sqinter>i\<in>A. f i) = x"
   120   unfolding INF_def by (rule Inf_eqI) auto
   121 
   122 lemma INF_lower: "i \<in> A \<Longrightarrow> (\<Sqinter>i\<in>A. f i) \<sqsubseteq> f i"
   123   by (auto simp add: INF_def intro: Inf_lower)
   124 
   125 lemma INF_greatest: "(\<And>i. i \<in> A \<Longrightarrow> u \<sqsubseteq> f i) \<Longrightarrow> u \<sqsubseteq> (\<Sqinter>i\<in>A. f i)"
   126   by (auto simp add: INF_def intro: Inf_greatest)
   127 
   128 lemma SUP_upper: "i \<in> A \<Longrightarrow> f i \<sqsubseteq> (\<Squnion>i\<in>A. f i)"
   129   by (auto simp add: SUP_def intro: Sup_upper)
   130 
   131 lemma SUP_least: "(\<And>i. i \<in> A \<Longrightarrow> f i \<sqsubseteq> u) \<Longrightarrow> (\<Squnion>i\<in>A. f i) \<sqsubseteq> u"
   132   by (auto simp add: SUP_def intro: Sup_least)
   133 
   134 lemma Inf_lower2: "u \<in> A \<Longrightarrow> u \<sqsubseteq> v \<Longrightarrow> \<Sqinter>A \<sqsubseteq> v"
   135   using Inf_lower [of u A] by auto
   136 
   137 lemma INF_lower2: "i \<in> A \<Longrightarrow> f i \<sqsubseteq> u \<Longrightarrow> (\<Sqinter>i\<in>A. f i) \<sqsubseteq> u"
   138   using INF_lower [of i A f] by auto
   139 
   140 lemma Sup_upper2: "u \<in> A \<Longrightarrow> v \<sqsubseteq> u \<Longrightarrow> v \<sqsubseteq> \<Squnion>A"
   141   using Sup_upper [of u A] by auto
   142 
   143 lemma SUP_upper2: "i \<in> A \<Longrightarrow> u \<sqsubseteq> f i \<Longrightarrow> u \<sqsubseteq> (\<Squnion>i\<in>A. f i)"
   144   using SUP_upper [of i A f] by auto
   145 
   146 lemma le_Inf_iff: "b \<sqsubseteq> \<Sqinter>A \<longleftrightarrow> (\<forall>a\<in>A. b \<sqsubseteq> a)"
   147   by (auto intro: Inf_greatest dest: Inf_lower)
   148 
   149 lemma le_INF_iff: "u \<sqsubseteq> (\<Sqinter>i\<in>A. f i) \<longleftrightarrow> (\<forall>i\<in>A. u \<sqsubseteq> f i)"
   150   by (auto simp add: INF_def le_Inf_iff)
   151 
   152 lemma Sup_le_iff: "\<Squnion>A \<sqsubseteq> b \<longleftrightarrow> (\<forall>a\<in>A. a \<sqsubseteq> b)"
   153   by (auto intro: Sup_least dest: Sup_upper)
   154 
   155 lemma SUP_le_iff: "(\<Squnion>i\<in>A. f i) \<sqsubseteq> u \<longleftrightarrow> (\<forall>i\<in>A. f i \<sqsubseteq> u)"
   156   by (auto simp add: SUP_def Sup_le_iff)
   157 
   158 lemma Inf_insert [simp]: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
   159   by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower)
   160 
   161 lemma INF_insert: "(\<Sqinter>x\<in>insert a A. f x) = f a \<sqinter> INFI A f"
   162   by (simp add: INF_def)
   163 
   164 lemma Sup_insert [simp]: "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
   165   by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper)
   166 
   167 lemma SUP_insert: "(\<Squnion>x\<in>insert a A. f x) = f a \<squnion> SUPR A f"
   168   by (simp add: SUP_def)
   169 
   170 lemma INF_empty [simp]: "(\<Sqinter>x\<in>{}. f x) = \<top>"
   171   by (simp add: INF_def)
   172 
   173 lemma SUP_empty [simp]: "(\<Squnion>x\<in>{}. f x) = \<bottom>"
   174   by (simp add: SUP_def)
   175 
   176 lemma Inf_UNIV [simp]:
   177   "\<Sqinter>UNIV = \<bottom>"
   178   by (auto intro!: antisym Inf_lower)
   179 
   180 lemma Sup_UNIV [simp]:
   181   "\<Squnion>UNIV = \<top>"
   182   by (auto intro!: antisym Sup_upper)
   183 
   184 lemma INF_image [simp]: "(\<Sqinter>x\<in>f`A. g x) = (\<Sqinter>x\<in>A. g (f x))"
   185   by (simp add: INF_def image_image)
   186 
   187 lemma SUP_image [simp]: "(\<Squnion>x\<in>f`A. g x) = (\<Squnion>x\<in>A. g (f x))"
   188   by (simp add: SUP_def image_image)
   189 
   190 lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<sqsubseteq> a}"
   191   by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
   192 
   193 lemma Sup_Inf:  "\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<sqsubseteq> b}"
   194   by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
   195 
   196 lemma Inf_superset_mono: "B \<subseteq> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> \<Sqinter>B"
   197   by (auto intro: Inf_greatest Inf_lower)
   198 
   199 lemma Sup_subset_mono: "A \<subseteq> B \<Longrightarrow> \<Squnion>A \<sqsubseteq> \<Squnion>B"
   200   by (auto intro: Sup_least Sup_upper)
   201 
   202 lemma INF_cong:
   203   "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Sqinter>x\<in>A. C x) = (\<Sqinter>x\<in>B. D x)"
   204   by (simp add: INF_def image_def)
   205 
   206 lemma SUP_cong:
   207   "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Squnion>x\<in>A. C x) = (\<Squnion>x\<in>B. D x)"
   208   by (simp add: SUP_def image_def)
   209 
   210 lemma Inf_mono:
   211   assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<sqsubseteq> b"
   212   shows "\<Sqinter>A \<sqsubseteq> \<Sqinter>B"
   213 proof (rule Inf_greatest)
   214   fix b assume "b \<in> B"
   215   with assms obtain a where "a \<in> A" and "a \<sqsubseteq> b" by blast
   216   from `a \<in> A` have "\<Sqinter>A \<sqsubseteq> a" by (rule Inf_lower)
   217   with `a \<sqsubseteq> b` show "\<Sqinter>A \<sqsubseteq> b" by auto
   218 qed
   219 
   220 lemma INF_mono:
   221   "(\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<sqsubseteq> g m) \<Longrightarrow> (\<Sqinter>n\<in>A. f n) \<sqsubseteq> (\<Sqinter>n\<in>B. g n)"
   222   unfolding INF_def by (rule Inf_mono) fast
   223 
   224 lemma Sup_mono:
   225   assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<sqsubseteq> b"
   226   shows "\<Squnion>A \<sqsubseteq> \<Squnion>B"
   227 proof (rule Sup_least)
   228   fix a assume "a \<in> A"
   229   with assms obtain b where "b \<in> B" and "a \<sqsubseteq> b" by blast
   230   from `b \<in> B` have "b \<sqsubseteq> \<Squnion>B" by (rule Sup_upper)
   231   with `a \<sqsubseteq> b` show "a \<sqsubseteq> \<Squnion>B" by auto
   232 qed
   233 
   234 lemma SUP_mono:
   235   "(\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<sqsubseteq> g m) \<Longrightarrow> (\<Squnion>n\<in>A. f n) \<sqsubseteq> (\<Squnion>n\<in>B. g n)"
   236   unfolding SUP_def by (rule Sup_mono) fast
   237 
   238 lemma INF_superset_mono:
   239   "B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<sqsubseteq> (\<Sqinter>x\<in>B. g x)"
   240   -- {* The last inclusion is POSITIVE! *}
   241   by (blast intro: INF_mono dest: subsetD)
   242 
   243 lemma SUP_subset_mono:
   244   "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<sqsubseteq> g x) \<Longrightarrow> (\<Squnion>x\<in>A. f x) \<sqsubseteq> (\<Squnion>x\<in>B. g x)"
   245   by (blast intro: SUP_mono dest: subsetD)
   246 
   247 lemma Inf_less_eq:
   248   assumes "\<And>v. v \<in> A \<Longrightarrow> v \<sqsubseteq> u"
   249     and "A \<noteq> {}"
   250   shows "\<Sqinter>A \<sqsubseteq> u"
   251 proof -
   252   from `A \<noteq> {}` obtain v where "v \<in> A" by blast
   253   moreover with assms have "v \<sqsubseteq> u" by blast
   254   ultimately show ?thesis by (rule Inf_lower2)
   255 qed
   256 
   257 lemma less_eq_Sup:
   258   assumes "\<And>v. v \<in> A \<Longrightarrow> u \<sqsubseteq> v"
   259     and "A \<noteq> {}"
   260   shows "u \<sqsubseteq> \<Squnion>A"
   261 proof -
   262   from `A \<noteq> {}` obtain v where "v \<in> A" by blast
   263   moreover with assms have "u \<sqsubseteq> v" by blast
   264   ultimately show ?thesis by (rule Sup_upper2)
   265 qed
   266 
   267 lemma SUPR_eq:
   268   assumes "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<le> g j"
   269   assumes "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<le> f i"
   270   shows "(SUP i:A. f i) = (SUP j:B. g j)"
   271   by (intro antisym SUP_least) (blast intro: SUP_upper2 dest: assms)+
   272 
   273 lemma INFI_eq:
   274   assumes "\<And>i. i \<in> A \<Longrightarrow> \<exists>j\<in>B. f i \<ge> g j"
   275   assumes "\<And>j. j \<in> B \<Longrightarrow> \<exists>i\<in>A. g j \<ge> f i"
   276   shows "(INF i:A. f i) = (INF j:B. g j)"
   277   by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+
   278 
   279 lemma less_eq_Inf_inter: "\<Sqinter>A \<squnion> \<Sqinter>B \<sqsubseteq> \<Sqinter>(A \<inter> B)"
   280   by (auto intro: Inf_greatest Inf_lower)
   281 
   282 lemma Sup_inter_less_eq: "\<Squnion>(A \<inter> B) \<sqsubseteq> \<Squnion>A \<sqinter> \<Squnion>B "
   283   by (auto intro: Sup_least Sup_upper)
   284 
   285 lemma Inf_union_distrib: "\<Sqinter>(A \<union> B) = \<Sqinter>A \<sqinter> \<Sqinter>B"
   286   by (rule antisym) (auto intro: Inf_greatest Inf_lower le_infI1 le_infI2)
   287 
   288 lemma INF_union:
   289   "(\<Sqinter>i \<in> A \<union> B. M i) = (\<Sqinter>i \<in> A. M i) \<sqinter> (\<Sqinter>i\<in>B. M i)"
   290   by (auto intro!: antisym INF_mono intro: le_infI1 le_infI2 INF_greatest INF_lower)
   291 
   292 lemma Sup_union_distrib: "\<Squnion>(A \<union> B) = \<Squnion>A \<squnion> \<Squnion>B"
   293   by (rule antisym) (auto intro: Sup_least Sup_upper le_supI1 le_supI2)
   294 
   295 lemma SUP_union:
   296   "(\<Squnion>i \<in> A \<union> B. M i) = (\<Squnion>i \<in> A. M i) \<squnion> (\<Squnion>i\<in>B. M i)"
   297   by (auto intro!: antisym SUP_mono intro: le_supI1 le_supI2 SUP_least SUP_upper)
   298 
   299 lemma INF_inf_distrib: "(\<Sqinter>a\<in>A. f a) \<sqinter> (\<Sqinter>a\<in>A. g a) = (\<Sqinter>a\<in>A. f a \<sqinter> g a)"
   300   by (rule antisym) (rule INF_greatest, auto intro: le_infI1 le_infI2 INF_lower INF_mono)
   301 
   302 lemma SUP_sup_distrib: "(\<Squnion>a\<in>A. f a) \<squnion> (\<Squnion>a\<in>A. g a) = (\<Squnion>a\<in>A. f a \<squnion> g a)" (is "?L = ?R")
   303 proof (rule antisym)
   304   show "?L \<le> ?R" by (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono)
   305 next
   306   show "?R \<le> ?L" by (rule SUP_least) (auto intro: le_supI1 le_supI2 SUP_upper)
   307 qed
   308 
   309 lemma Inf_top_conv [simp, no_atp]:
   310   "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
   311   "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
   312 proof -
   313   show "\<Sqinter>A = \<top> \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)"
   314   proof
   315     assume "\<forall>x\<in>A. x = \<top>"
   316     then have "A = {} \<or> A = {\<top>}" by auto
   317     then show "\<Sqinter>A = \<top>" by auto
   318   next
   319     assume "\<Sqinter>A = \<top>"
   320     show "\<forall>x\<in>A. x = \<top>"
   321     proof (rule ccontr)
   322       assume "\<not> (\<forall>x\<in>A. x = \<top>)"
   323       then obtain x where "x \<in> A" and "x \<noteq> \<top>" by blast
   324       then obtain B where "A = insert x B" by blast
   325       with `\<Sqinter>A = \<top>` `x \<noteq> \<top>` show False by simp
   326     qed
   327   qed
   328   then show "\<top> = \<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<top>)" by auto
   329 qed
   330 
   331 lemma INF_top_conv [simp]:
   332  "(\<Sqinter>x\<in>A. B x) = \<top> \<longleftrightarrow> (\<forall>x\<in>A. B x = \<top>)"
   333  "\<top> = (\<Sqinter>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = \<top>)"
   334   by (auto simp add: INF_def)
   335 
   336 lemma Sup_bot_conv [simp, no_atp]:
   337   "\<Squnion>A = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?P)
   338   "\<bottom> = \<Squnion>A \<longleftrightarrow> (\<forall>x\<in>A. x = \<bottom>)" (is ?Q)
   339   using dual_complete_lattice
   340   by (rule complete_lattice.Inf_top_conv)+
   341 
   342 lemma SUP_bot_conv [simp]:
   343  "(\<Squnion>x\<in>A. B x) = \<bottom> \<longleftrightarrow> (\<forall>x\<in>A. B x = \<bottom>)"
   344  "\<bottom> = (\<Squnion>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = \<bottom>)"
   345   by (auto simp add: SUP_def)
   346 
   347 lemma INF_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Sqinter>i\<in>A. f) = f"
   348   by (auto intro: antisym INF_lower INF_greatest)
   349 
   350 lemma SUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Squnion>i\<in>A. f) = f"
   351   by (auto intro: antisym SUP_upper SUP_least)
   352 
   353 lemma INF_top [simp]: "(\<Sqinter>x\<in>A. \<top>) = \<top>"
   354   by (cases "A = {}") simp_all
   355 
   356 lemma SUP_bot [simp]: "(\<Squnion>x\<in>A. \<bottom>) = \<bottom>"
   357   by (cases "A = {}") simp_all
   358 
   359 lemma INF_commute: "(\<Sqinter>i\<in>A. \<Sqinter>j\<in>B. f i j) = (\<Sqinter>j\<in>B. \<Sqinter>i\<in>A. f i j)"
   360   by (iprover intro: INF_lower INF_greatest order_trans antisym)
   361 
   362 lemma SUP_commute: "(\<Squnion>i\<in>A. \<Squnion>j\<in>B. f i j) = (\<Squnion>j\<in>B. \<Squnion>i\<in>A. f i j)"
   363   by (iprover intro: SUP_upper SUP_least order_trans antisym)
   364 
   365 lemma INF_absorb:
   366   assumes "k \<in> I"
   367   shows "A k \<sqinter> (\<Sqinter>i\<in>I. A i) = (\<Sqinter>i\<in>I. A i)"
   368 proof -
   369   from assms obtain J where "I = insert k J" by blast
   370   then show ?thesis by (simp add: INF_insert)
   371 qed
   372 
   373 lemma SUP_absorb:
   374   assumes "k \<in> I"
   375   shows "A k \<squnion> (\<Squnion>i\<in>I. A i) = (\<Squnion>i\<in>I. A i)"
   376 proof -
   377   from assms obtain J where "I = insert k J" by blast
   378   then show ?thesis by (simp add: SUP_insert)
   379 qed
   380 
   381 lemma INF_constant:
   382   "(\<Sqinter>y\<in>A. c) = (if A = {} then \<top> else c)"
   383   by simp
   384 
   385 lemma SUP_constant:
   386   "(\<Squnion>y\<in>A. c) = (if A = {} then \<bottom> else c)"
   387   by simp
   388 
   389 lemma less_INF_D:
   390   assumes "y < (\<Sqinter>i\<in>A. f i)" "i \<in> A" shows "y < f i"
   391 proof -
   392   note `y < (\<Sqinter>i\<in>A. f i)`
   393   also have "(\<Sqinter>i\<in>A. f i) \<le> f i" using `i \<in> A`
   394     by (rule INF_lower)
   395   finally show "y < f i" .
   396 qed
   397 
   398 lemma SUP_lessD:
   399   assumes "(\<Squnion>i\<in>A. f i) < y" "i \<in> A" shows "f i < y"
   400 proof -
   401   have "f i \<le> (\<Squnion>i\<in>A. f i)" using `i \<in> A`
   402     by (rule SUP_upper)
   403   also note `(\<Squnion>i\<in>A. f i) < y`
   404   finally show "f i < y" .
   405 qed
   406 
   407 lemma INF_UNIV_bool_expand:
   408   "(\<Sqinter>b. A b) = A True \<sqinter> A False"
   409   by (simp add: UNIV_bool INF_insert inf_commute)
   410 
   411 lemma SUP_UNIV_bool_expand:
   412   "(\<Squnion>b. A b) = A True \<squnion> A False"
   413   by (simp add: UNIV_bool SUP_insert sup_commute)
   414 
   415 lemma Inf_le_Sup: "A \<noteq> {} \<Longrightarrow> Inf A \<le> Sup A"
   416   by (blast intro: Sup_upper2 Inf_lower ex_in_conv)
   417 
   418 lemma INF_le_SUP: "A \<noteq> {} \<Longrightarrow> INFI A f \<le> SUPR A f"
   419   unfolding INF_def SUP_def by (rule Inf_le_Sup) auto
   420 
   421 end
   422 
   423 class complete_distrib_lattice = complete_lattice +
   424   assumes sup_Inf: "a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)"
   425   assumes inf_Sup: "a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
   426 begin
   427 
   428 lemma sup_INF:
   429   "a \<squnion> (\<Sqinter>b\<in>B. f b) = (\<Sqinter>b\<in>B. a \<squnion> f b)"
   430   by (simp add: INF_def sup_Inf image_image)
   431 
   432 lemma inf_SUP:
   433   "a \<sqinter> (\<Squnion>b\<in>B. f b) = (\<Squnion>b\<in>B. a \<sqinter> f b)"
   434   by (simp add: SUP_def inf_Sup image_image)
   435 
   436 lemma dual_complete_distrib_lattice:
   437   "class.complete_distrib_lattice Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom>"
   438   apply (rule class.complete_distrib_lattice.intro)
   439   apply (fact dual_complete_lattice)
   440   apply (rule class.complete_distrib_lattice_axioms.intro)
   441   apply (simp_all only: INF_foundation_dual SUP_foundation_dual inf_Sup sup_Inf)
   442   done
   443 
   444 subclass distrib_lattice proof
   445   fix a b c
   446   from sup_Inf have "a \<squnion> \<Sqinter>{b, c} = (\<Sqinter>d\<in>{b, c}. a \<squnion> d)" .
   447   then show "a \<squnion> b \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c)" by (simp add: INF_def)
   448 qed
   449 
   450 lemma Inf_sup:
   451   "\<Sqinter>B \<squnion> a = (\<Sqinter>b\<in>B. b \<squnion> a)"
   452   by (simp add: sup_Inf sup_commute)
   453 
   454 lemma Sup_inf:
   455   "\<Squnion>B \<sqinter> a = (\<Squnion>b\<in>B. b \<sqinter> a)"
   456   by (simp add: inf_Sup inf_commute)
   457 
   458 lemma INF_sup: 
   459   "(\<Sqinter>b\<in>B. f b) \<squnion> a = (\<Sqinter>b\<in>B. f b \<squnion> a)"
   460   by (simp add: sup_INF sup_commute)
   461 
   462 lemma SUP_inf:
   463   "(\<Squnion>b\<in>B. f b) \<sqinter> a = (\<Squnion>b\<in>B. f b \<sqinter> a)"
   464   by (simp add: inf_SUP inf_commute)
   465 
   466 lemma Inf_sup_eq_top_iff:
   467   "(\<Sqinter>B \<squnion> a = \<top>) \<longleftrightarrow> (\<forall>b\<in>B. b \<squnion> a = \<top>)"
   468   by (simp only: Inf_sup INF_top_conv)
   469 
   470 lemma Sup_inf_eq_bot_iff:
   471   "(\<Squnion>B \<sqinter> a = \<bottom>) \<longleftrightarrow> (\<forall>b\<in>B. b \<sqinter> a = \<bottom>)"
   472   by (simp only: Sup_inf SUP_bot_conv)
   473 
   474 lemma INF_sup_distrib2:
   475   "(\<Sqinter>a\<in>A. f a) \<squnion> (\<Sqinter>b\<in>B. g b) = (\<Sqinter>a\<in>A. \<Sqinter>b\<in>B. f a \<squnion> g b)"
   476   by (subst INF_commute) (simp add: sup_INF INF_sup)
   477 
   478 lemma SUP_inf_distrib2:
   479   "(\<Squnion>a\<in>A. f a) \<sqinter> (\<Squnion>b\<in>B. g b) = (\<Squnion>a\<in>A. \<Squnion>b\<in>B. f a \<sqinter> g b)"
   480   by (subst SUP_commute) (simp add: inf_SUP SUP_inf)
   481 
   482 end
   483 
   484 class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice
   485 begin
   486 
   487 lemma dual_complete_boolean_algebra:
   488   "class.complete_boolean_algebra Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom> (\<lambda>x y. x \<squnion> - y) uminus"
   489   by (rule class.complete_boolean_algebra.intro, rule dual_complete_distrib_lattice, rule dual_boolean_algebra)
   490 
   491 lemma uminus_Inf:
   492   "- (\<Sqinter>A) = \<Squnion>(uminus ` A)"
   493 proof (rule antisym)
   494   show "- \<Sqinter>A \<le> \<Squnion>(uminus ` A)"
   495     by (rule compl_le_swap2, rule Inf_greatest, rule compl_le_swap2, rule Sup_upper) simp
   496   show "\<Squnion>(uminus ` A) \<le> - \<Sqinter>A"
   497     by (rule Sup_least, rule compl_le_swap1, rule Inf_lower) auto
   498 qed
   499 
   500 lemma uminus_INF: "- (\<Sqinter>x\<in>A. B x) = (\<Squnion>x\<in>A. - B x)"
   501   by (simp add: INF_def SUP_def uminus_Inf image_image)
   502 
   503 lemma uminus_Sup:
   504   "- (\<Squnion>A) = \<Sqinter>(uminus ` A)"
   505 proof -
   506   have "\<Squnion>A = - \<Sqinter>(uminus ` A)" by (simp add: image_image uminus_Inf)
   507   then show ?thesis by simp
   508 qed
   509   
   510 lemma uminus_SUP: "- (\<Squnion>x\<in>A. B x) = (\<Sqinter>x\<in>A. - B x)"
   511   by (simp add: INF_def SUP_def uminus_Sup image_image)
   512 
   513 end
   514 
   515 class complete_linorder = linorder + complete_lattice
   516 begin
   517 
   518 lemma dual_complete_linorder:
   519   "class.complete_linorder Sup Inf sup (op \<ge>) (op >) inf \<top> \<bottom>"
   520   by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder)
   521 
   522 lemma complete_linorder_inf_min: "inf = min"
   523   by (auto intro: antisym simp add: min_def fun_eq_iff)
   524 
   525 lemma complete_linorder_sup_max: "sup = max"
   526   by (auto intro: antisym simp add: max_def fun_eq_iff)
   527 
   528 lemma Inf_less_iff:
   529   "\<Sqinter>S \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>S. x \<sqsubset> a)"
   530   unfolding not_le [symmetric] le_Inf_iff by auto
   531 
   532 lemma INF_less_iff:
   533   "(\<Sqinter>i\<in>A. f i) \<sqsubset> a \<longleftrightarrow> (\<exists>x\<in>A. f x \<sqsubset> a)"
   534   unfolding INF_def Inf_less_iff by auto
   535 
   536 lemma less_Sup_iff:
   537   "a \<sqsubset> \<Squnion>S \<longleftrightarrow> (\<exists>x\<in>S. a \<sqsubset> x)"
   538   unfolding not_le [symmetric] Sup_le_iff by auto
   539 
   540 lemma less_SUP_iff:
   541   "a \<sqsubset> (\<Squnion>i\<in>A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a \<sqsubset> f x)"
   542   unfolding SUP_def less_Sup_iff by auto
   543 
   544 lemma Sup_eq_top_iff [simp]:
   545   "\<Squnion>A = \<top> \<longleftrightarrow> (\<forall>x<\<top>. \<exists>i\<in>A. x < i)"
   546 proof
   547   assume *: "\<Squnion>A = \<top>"
   548   show "(\<forall>x<\<top>. \<exists>i\<in>A. x < i)" unfolding * [symmetric]
   549   proof (intro allI impI)
   550     fix x assume "x < \<Squnion>A" then show "\<exists>i\<in>A. x < i"
   551       unfolding less_Sup_iff by auto
   552   qed
   553 next
   554   assume *: "\<forall>x<\<top>. \<exists>i\<in>A. x < i"
   555   show "\<Squnion>A = \<top>"
   556   proof (rule ccontr)
   557     assume "\<Squnion>A \<noteq> \<top>"
   558     with top_greatest [of "\<Squnion>A"]
   559     have "\<Squnion>A < \<top>" unfolding le_less by auto
   560     then have "\<Squnion>A < \<Squnion>A"
   561       using * unfolding less_Sup_iff by auto
   562     then show False by auto
   563   qed
   564 qed
   565 
   566 lemma SUP_eq_top_iff [simp]:
   567   "(\<Squnion>i\<in>A. f i) = \<top> \<longleftrightarrow> (\<forall>x<\<top>. \<exists>i\<in>A. x < f i)"
   568   unfolding SUP_def by auto
   569 
   570 lemma Inf_eq_bot_iff [simp]:
   571   "\<Sqinter>A = \<bottom> \<longleftrightarrow> (\<forall>x>\<bottom>. \<exists>i\<in>A. i < x)"
   572   using dual_complete_linorder
   573   by (rule complete_linorder.Sup_eq_top_iff)
   574 
   575 lemma INF_eq_bot_iff [simp]:
   576   "(\<Sqinter>i\<in>A. f i) = \<bottom> \<longleftrightarrow> (\<forall>x>\<bottom>. \<exists>i\<in>A. f i < x)"
   577   unfolding INF_def by auto
   578 
   579 lemma le_Sup_iff: "x \<le> \<Squnion>A \<longleftrightarrow> (\<forall>y<x. \<exists>a\<in>A. y < a)"
   580 proof safe
   581   fix y assume "x \<le> \<Squnion>A" "y < x"
   582   then have "y < \<Squnion>A" by auto
   583   then show "\<exists>a\<in>A. y < a"
   584     unfolding less_Sup_iff .
   585 qed (auto elim!: allE[of _ "\<Squnion>A"] simp add: not_le[symmetric] Sup_upper)
   586 
   587 lemma le_SUP_iff: "x \<le> SUPR A f \<longleftrightarrow> (\<forall>y<x. \<exists>i\<in>A. y < f i)"
   588   unfolding le_Sup_iff SUP_def by simp
   589 
   590 lemma Inf_le_iff: "\<Sqinter>A \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>a\<in>A. y > a)"
   591 proof safe
   592   fix y assume "x \<ge> \<Sqinter>A" "y > x"
   593   then have "y > \<Sqinter>A" by auto
   594   then show "\<exists>a\<in>A. y > a"
   595     unfolding Inf_less_iff .
   596 qed (auto elim!: allE[of _ "\<Sqinter>A"] simp add: not_le[symmetric] Inf_lower)
   597 
   598 lemma INF_le_iff:
   599   "INFI A f \<le> x \<longleftrightarrow> (\<forall>y>x. \<exists>i\<in>A. y > f i)"
   600   unfolding Inf_le_iff INF_def by simp
   601 
   602 subclass complete_distrib_lattice
   603 proof
   604   fix a and B
   605   show "a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)" and "a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
   606     by (safe intro!: INF_eqI [symmetric] sup_mono Inf_lower SUP_eqI [symmetric] inf_mono Sup_upper)
   607       (auto simp: not_less [symmetric] Inf_less_iff less_Sup_iff
   608         le_max_iff_disj complete_linorder_sup_max min_le_iff_disj complete_linorder_inf_min)
   609 qed
   610 
   611 end
   612 
   613 
   614 subsection {* Complete lattice on @{typ bool} *}
   615 
   616 instantiation bool :: complete_lattice
   617 begin
   618 
   619 definition
   620   [simp, code]: "\<Sqinter>A \<longleftrightarrow> False \<notin> A"
   621 
   622 definition
   623   [simp, code]: "\<Squnion>A \<longleftrightarrow> True \<in> A"
   624 
   625 instance proof
   626 qed (auto intro: bool_induct)
   627 
   628 end
   629 
   630 lemma not_False_in_image_Ball [simp]:
   631   "False \<notin> P ` A \<longleftrightarrow> Ball A P"
   632   by auto
   633 
   634 lemma True_in_image_Bex [simp]:
   635   "True \<in> P ` A \<longleftrightarrow> Bex A P"
   636   by auto
   637 
   638 lemma INF_bool_eq [simp]:
   639   "INFI = Ball"
   640   by (simp add: fun_eq_iff INF_def)
   641 
   642 lemma SUP_bool_eq [simp]:
   643   "SUPR = Bex"
   644   by (simp add: fun_eq_iff SUP_def)
   645 
   646 instance bool :: complete_boolean_algebra proof
   647 qed (auto intro: bool_induct)
   648 
   649 
   650 subsection {* Complete lattice on @{typ "_ \<Rightarrow> _"} *}
   651 
   652 instantiation "fun" :: (type, complete_lattice) complete_lattice
   653 begin
   654 
   655 definition
   656   "\<Sqinter>A = (\<lambda>x. \<Sqinter>f\<in>A. f x)"
   657 
   658 lemma Inf_apply [simp, code]:
   659   "(\<Sqinter>A) x = (\<Sqinter>f\<in>A. f x)"
   660   by (simp add: Inf_fun_def)
   661 
   662 definition
   663   "\<Squnion>A = (\<lambda>x. \<Squnion>f\<in>A. f x)"
   664 
   665 lemma Sup_apply [simp, code]:
   666   "(\<Squnion>A) x = (\<Squnion>f\<in>A. f x)"
   667   by (simp add: Sup_fun_def)
   668 
   669 instance proof
   670 qed (auto simp add: le_fun_def intro: INF_lower INF_greatest SUP_upper SUP_least)
   671 
   672 end
   673 
   674 lemma INF_apply [simp]:
   675   "(\<Sqinter>y\<in>A. f y) x = (\<Sqinter>y\<in>A. f y x)"
   676   by (auto intro: arg_cong [of _ _ Inf] simp add: INF_def)
   677 
   678 lemma SUP_apply [simp]:
   679   "(\<Squnion>y\<in>A. f y) x = (\<Squnion>y\<in>A. f y x)"
   680   by (auto intro: arg_cong [of _ _ Sup] simp add: SUP_def)
   681 
   682 instance "fun" :: (type, complete_distrib_lattice) complete_distrib_lattice proof
   683 qed (auto simp add: INF_def SUP_def inf_Sup sup_Inf image_image)
   684 
   685 instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra ..
   686 
   687 
   688 subsection {* Complete lattice on unary and binary predicates *}
   689 
   690 lemma INF1_iff: "(\<Sqinter>x\<in>A. B x) b = (\<forall>x\<in>A. B x b)"
   691   by simp
   692 
   693 lemma INF2_iff: "(\<Sqinter>x\<in>A. B x) b c = (\<forall>x\<in>A. B x b c)"
   694   by simp
   695 
   696 lemma INF1_I: "(\<And>x. x \<in> A \<Longrightarrow> B x b) \<Longrightarrow> (\<Sqinter>x\<in>A. B x) b"
   697   by auto
   698 
   699 lemma INF2_I: "(\<And>x. x \<in> A \<Longrightarrow> B x b c) \<Longrightarrow> (\<Sqinter>x\<in>A. B x) b c"
   700   by auto
   701 
   702 lemma INF1_D: "(\<Sqinter>x\<in>A. B x) b \<Longrightarrow> a \<in> A \<Longrightarrow> B a b"
   703   by auto
   704 
   705 lemma INF2_D: "(\<Sqinter>x\<in>A. B x) b c \<Longrightarrow> a \<in> A \<Longrightarrow> B a b c"
   706   by auto
   707 
   708 lemma INF1_E: "(\<Sqinter>x\<in>A. B x) b \<Longrightarrow> (B a b \<Longrightarrow> R) \<Longrightarrow> (a \<notin> A \<Longrightarrow> R) \<Longrightarrow> R"
   709   by auto
   710 
   711 lemma INF2_E: "(\<Sqinter>x\<in>A. B x) b c \<Longrightarrow> (B a b c \<Longrightarrow> R) \<Longrightarrow> (a \<notin> A \<Longrightarrow> R) \<Longrightarrow> R"
   712   by auto
   713 
   714 lemma SUP1_iff: "(\<Squnion>x\<in>A. B x) b = (\<exists>x\<in>A. B x b)"
   715   by simp
   716 
   717 lemma SUP2_iff: "(\<Squnion>x\<in>A. B x) b c = (\<exists>x\<in>A. B x b c)"
   718   by simp
   719 
   720 lemma SUP1_I: "a \<in> A \<Longrightarrow> B a b \<Longrightarrow> (\<Squnion>x\<in>A. B x) b"
   721   by auto
   722 
   723 lemma SUP2_I: "a \<in> A \<Longrightarrow> B a b c \<Longrightarrow> (\<Squnion>x\<in>A. B x) b c"
   724   by auto
   725 
   726 lemma SUP1_E: "(\<Squnion>x\<in>A. B x) b \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> B x b \<Longrightarrow> R) \<Longrightarrow> R"
   727   by auto
   728 
   729 lemma SUP2_E: "(\<Squnion>x\<in>A. B x) b c \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> B x b c \<Longrightarrow> R) \<Longrightarrow> R"
   730   by auto
   731 
   732 
   733 subsection {* Complete lattice on @{typ "_ set"} *}
   734 
   735 instantiation "set" :: (type) complete_lattice
   736 begin
   737 
   738 definition
   739   "\<Sqinter>A = {x. \<Sqinter>((\<lambda>B. x \<in> B) ` A)}"
   740 
   741 definition
   742   "\<Squnion>A = {x. \<Squnion>((\<lambda>B. x \<in> B) ` A)}"
   743 
   744 instance proof
   745 qed (auto simp add: less_eq_set_def Inf_set_def Sup_set_def le_fun_def)
   746 
   747 end
   748 
   749 instance "set" :: (type) complete_boolean_algebra
   750 proof
   751 qed (auto simp add: INF_def SUP_def Inf_set_def Sup_set_def image_def)
   752   
   753 
   754 subsubsection {* Inter *}
   755 
   756 abbreviation Inter :: "'a set set \<Rightarrow> 'a set" where
   757   "Inter S \<equiv> \<Sqinter>S"
   758   
   759 notation (xsymbols)
   760   Inter  ("\<Inter>_" [900] 900)
   761 
   762 lemma Inter_eq:
   763   "\<Inter>A = {x. \<forall>B \<in> A. x \<in> B}"
   764 proof (rule set_eqI)
   765   fix x
   766   have "(\<forall>Q\<in>{P. \<exists>B\<in>A. P \<longleftrightarrow> x \<in> B}. Q) \<longleftrightarrow> (\<forall>B\<in>A. x \<in> B)"
   767     by auto
   768   then show "x \<in> \<Inter>A \<longleftrightarrow> x \<in> {x. \<forall>B \<in> A. x \<in> B}"
   769     by (simp add: Inf_set_def image_def)
   770 qed
   771 
   772 lemma Inter_iff [simp,no_atp]: "A \<in> \<Inter>C \<longleftrightarrow> (\<forall>X\<in>C. A \<in> X)"
   773   by (unfold Inter_eq) blast
   774 
   775 lemma InterI [intro!]: "(\<And>X. X \<in> C \<Longrightarrow> A \<in> X) \<Longrightarrow> A \<in> \<Inter>C"
   776   by (simp add: Inter_eq)
   777 
   778 text {*
   779   \medskip A ``destruct'' rule -- every @{term X} in @{term C}
   780   contains @{term A} as an element, but @{prop "A \<in> X"} can hold when
   781   @{prop "X \<in> C"} does not!  This rule is analogous to @{text spec}.
   782 *}
   783 
   784 lemma InterD [elim, Pure.elim]: "A \<in> \<Inter>C \<Longrightarrow> X \<in> C \<Longrightarrow> A \<in> X"
   785   by auto
   786 
   787 lemma InterE [elim]: "A \<in> \<Inter>C \<Longrightarrow> (X \<notin> C \<Longrightarrow> R) \<Longrightarrow> (A \<in> X \<Longrightarrow> R) \<Longrightarrow> R"
   788   -- {* ``Classical'' elimination rule -- does not require proving
   789     @{prop "X \<in> C"}. *}
   790   by (unfold Inter_eq) blast
   791 
   792 lemma Inter_lower: "B \<in> A \<Longrightarrow> \<Inter>A \<subseteq> B"
   793   by (fact Inf_lower)
   794 
   795 lemma Inter_subset:
   796   "(\<And>X. X \<in> A \<Longrightarrow> X \<subseteq> B) \<Longrightarrow> A \<noteq> {} \<Longrightarrow> \<Inter>A \<subseteq> B"
   797   by (fact Inf_less_eq)
   798 
   799 lemma Inter_greatest: "(\<And>X. X \<in> A \<Longrightarrow> C \<subseteq> X) \<Longrightarrow> C \<subseteq> Inter A"
   800   by (fact Inf_greatest)
   801 
   802 lemma Inter_empty: "\<Inter>{} = UNIV"
   803   by (fact Inf_empty) (* already simp *)
   804 
   805 lemma Inter_UNIV: "\<Inter>UNIV = {}"
   806   by (fact Inf_UNIV) (* already simp *)
   807 
   808 lemma Inter_insert: "\<Inter>(insert a B) = a \<inter> \<Inter>B"
   809   by (fact Inf_insert) (* already simp *)
   810 
   811 lemma Inter_Un_subset: "\<Inter>A \<union> \<Inter>B \<subseteq> \<Inter>(A \<inter> B)"
   812   by (fact less_eq_Inf_inter)
   813 
   814 lemma Inter_Un_distrib: "\<Inter>(A \<union> B) = \<Inter>A \<inter> \<Inter>B"
   815   by (fact Inf_union_distrib)
   816 
   817 lemma Inter_UNIV_conv [simp, no_atp]:
   818   "\<Inter>A = UNIV \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
   819   "UNIV = \<Inter>A \<longleftrightarrow> (\<forall>x\<in>A. x = UNIV)"
   820   by (fact Inf_top_conv)+
   821 
   822 lemma Inter_anti_mono: "B \<subseteq> A \<Longrightarrow> \<Inter>A \<subseteq> \<Inter>B"
   823   by (fact Inf_superset_mono)
   824 
   825 
   826 subsubsection {* Intersections of families *}
   827 
   828 abbreviation INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
   829   "INTER \<equiv> INFI"
   830 
   831 text {*
   832   Note: must use name @{const INTER} here instead of @{text INT}
   833   to allow the following syntax coexist with the plain constant name.
   834 *}
   835 
   836 syntax
   837   "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3INT _./ _)" [0, 10] 10)
   838   "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3INT _:_./ _)" [0, 0, 10] 10)
   839 
   840 syntax (xsymbols)
   841   "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>_./ _)" [0, 10] 10)
   842   "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>_\<in>_./ _)" [0, 0, 10] 10)
   843 
   844 syntax (latex output)
   845   "_INTER1"     :: "pttrns => 'b set => 'b set"           ("(3\<Inter>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
   846   "_INTER"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
   847 
   848 translations
   849   "INT x y. B"  == "INT x. INT y. B"
   850   "INT x. B"    == "CONST INTER CONST UNIV (%x. B)"
   851   "INT x. B"    == "INT x:CONST UNIV. B"
   852   "INT x:A. B"  == "CONST INTER A (%x. B)"
   853 
   854 print_translation {*
   855   [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}]
   856 *} -- {* to avoid eta-contraction of body *}
   857 
   858 lemma INTER_eq:
   859   "(\<Inter>x\<in>A. B x) = {y. \<forall>x\<in>A. y \<in> B x}"
   860   by (auto simp add: INF_def)
   861 
   862 lemma Inter_image_eq [simp]:
   863   "\<Inter>(B`A) = (\<Inter>x\<in>A. B x)"
   864   by (rule sym) (fact INF_def)
   865 
   866 lemma INT_iff [simp]: "b \<in> (\<Inter>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. b \<in> B x)"
   867   by (auto simp add: INF_def image_def)
   868 
   869 lemma INT_I [intro!]: "(\<And>x. x \<in> A \<Longrightarrow> b \<in> B x) \<Longrightarrow> b \<in> (\<Inter>x\<in>A. B x)"
   870   by (auto simp add: INF_def image_def)
   871 
   872 lemma INT_D [elim, Pure.elim]: "b \<in> (\<Inter>x\<in>A. B x) \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> B a"
   873   by auto
   874 
   875 lemma INT_E [elim]: "b \<in> (\<Inter>x\<in>A. B x) \<Longrightarrow> (b \<in> B a \<Longrightarrow> R) \<Longrightarrow> (a \<notin> A \<Longrightarrow> R) \<Longrightarrow> R"
   876   -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a\<in>A"}. *}
   877   by (auto simp add: INF_def image_def)
   878 
   879 lemma INT_cong [cong]:
   880   "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Inter>x\<in>A. C x) = (\<Inter>x\<in>B. D x)"
   881   by (fact INF_cong)
   882 
   883 lemma Collect_ball_eq: "{x. \<forall>y\<in>A. P x y} = (\<Inter>y\<in>A. {x. P x y})"
   884   by blast
   885 
   886 lemma Collect_all_eq: "{x. \<forall>y. P x y} = (\<Inter>y. {x. P x y})"
   887   by blast
   888 
   889 lemma INT_lower: "a \<in> A \<Longrightarrow> (\<Inter>x\<in>A. B x) \<subseteq> B a"
   890   by (fact INF_lower)
   891 
   892 lemma INT_greatest: "(\<And>x. x \<in> A \<Longrightarrow> C \<subseteq> B x) \<Longrightarrow> C \<subseteq> (\<Inter>x\<in>A. B x)"
   893   by (fact INF_greatest)
   894 
   895 lemma INT_empty: "(\<Inter>x\<in>{}. B x) = UNIV"
   896   by (fact INF_empty)
   897 
   898 lemma INT_absorb: "k \<in> I \<Longrightarrow> A k \<inter> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. A i)"
   899   by (fact INF_absorb)
   900 
   901 lemma INT_subset_iff: "B \<subseteq> (\<Inter>i\<in>I. A i) \<longleftrightarrow> (\<forall>i\<in>I. B \<subseteq> A i)"
   902   by (fact le_INF_iff)
   903 
   904 lemma INT_insert [simp]: "(\<Inter>x \<in> insert a A. B x) = B a \<inter> INTER A B"
   905   by (fact INF_insert)
   906 
   907 lemma INT_Un: "(\<Inter>i \<in> A \<union> B. M i) = (\<Inter>i \<in> A. M i) \<inter> (\<Inter>i\<in>B. M i)"
   908   by (fact INF_union)
   909 
   910 lemma INT_insert_distrib:
   911   "u \<in> A \<Longrightarrow> (\<Inter>x\<in>A. insert a (B x)) = insert a (\<Inter>x\<in>A. B x)"
   912   by blast
   913 
   914 lemma INT_constant [simp]: "(\<Inter>y\<in>A. c) = (if A = {} then UNIV else c)"
   915   by (fact INF_constant)
   916 
   917 lemma INTER_UNIV_conv:
   918  "(UNIV = (\<Inter>x\<in>A. B x)) = (\<forall>x\<in>A. B x = UNIV)"
   919  "((\<Inter>x\<in>A. B x) = UNIV) = (\<forall>x\<in>A. B x = UNIV)"
   920   by (fact INF_top_conv)+ (* already simp *)
   921 
   922 lemma INT_bool_eq: "(\<Inter>b. A b) = A True \<inter> A False"
   923   by (fact INF_UNIV_bool_expand)
   924 
   925 lemma INT_anti_mono:
   926   "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<subseteq> g x) \<Longrightarrow> (\<Inter>x\<in>B. f x) \<subseteq> (\<Inter>x\<in>A. g x)"
   927   -- {* The last inclusion is POSITIVE! *}
   928   by (fact INF_superset_mono)
   929 
   930 lemma Pow_INT_eq: "Pow (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. Pow (B x))"
   931   by blast
   932 
   933 lemma vimage_INT: "f -` (\<Inter>x\<in>A. B x) = (\<Inter>x\<in>A. f -` B x)"
   934   by blast
   935 
   936 
   937 subsubsection {* Union *}
   938 
   939 abbreviation Union :: "'a set set \<Rightarrow> 'a set" where
   940   "Union S \<equiv> \<Squnion>S"
   941 
   942 notation (xsymbols)
   943   Union  ("\<Union>_" [900] 900)
   944 
   945 lemma Union_eq:
   946   "\<Union>A = {x. \<exists>B \<in> A. x \<in> B}"
   947 proof (rule set_eqI)
   948   fix x
   949   have "(\<exists>Q\<in>{P. \<exists>B\<in>A. P \<longleftrightarrow> x \<in> B}. Q) \<longleftrightarrow> (\<exists>B\<in>A. x \<in> B)"
   950     by auto
   951   then show "x \<in> \<Union>A \<longleftrightarrow> x \<in> {x. \<exists>B\<in>A. x \<in> B}"
   952     by (simp add: Sup_set_def image_def)
   953 qed
   954 
   955 lemma Union_iff [simp, no_atp]:
   956   "A \<in> \<Union>C \<longleftrightarrow> (\<exists>X\<in>C. A\<in>X)"
   957   by (unfold Union_eq) blast
   958 
   959 lemma UnionI [intro]:
   960   "X \<in> C \<Longrightarrow> A \<in> X \<Longrightarrow> A \<in> \<Union>C"
   961   -- {* The order of the premises presupposes that @{term C} is rigid;
   962     @{term A} may be flexible. *}
   963   by auto
   964 
   965 lemma UnionE [elim!]:
   966   "A \<in> \<Union>C \<Longrightarrow> (\<And>X. A \<in> X \<Longrightarrow> X \<in> C \<Longrightarrow> R) \<Longrightarrow> R"
   967   by auto
   968 
   969 lemma Union_upper: "B \<in> A \<Longrightarrow> B \<subseteq> \<Union>A"
   970   by (fact Sup_upper)
   971 
   972 lemma Union_least: "(\<And>X. X \<in> A \<Longrightarrow> X \<subseteq> C) \<Longrightarrow> \<Union>A \<subseteq> C"
   973   by (fact Sup_least)
   974 
   975 lemma Union_empty: "\<Union>{} = {}"
   976   by (fact Sup_empty) (* already simp *)
   977 
   978 lemma Union_UNIV: "\<Union>UNIV = UNIV"
   979   by (fact Sup_UNIV) (* already simp *)
   980 
   981 lemma Union_insert: "\<Union>insert a B = a \<union> \<Union>B"
   982   by (fact Sup_insert) (* already simp *)
   983 
   984 lemma Union_Un_distrib [simp]: "\<Union>(A \<union> B) = \<Union>A \<union> \<Union>B"
   985   by (fact Sup_union_distrib)
   986 
   987 lemma Union_Int_subset: "\<Union>(A \<inter> B) \<subseteq> \<Union>A \<inter> \<Union>B"
   988   by (fact Sup_inter_less_eq)
   989 
   990 lemma Union_empty_conv [no_atp]: "(\<Union>A = {}) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
   991   by (fact Sup_bot_conv) (* already simp *)
   992 
   993 lemma empty_Union_conv [no_atp]: "({} = \<Union>A) \<longleftrightarrow> (\<forall>x\<in>A. x = {})"
   994   by (fact Sup_bot_conv) (* already simp *)
   995 
   996 lemma subset_Pow_Union: "A \<subseteq> Pow (\<Union>A)"
   997   by blast
   998 
   999 lemma Union_Pow_eq [simp]: "\<Union>(Pow A) = A"
  1000   by blast
  1001 
  1002 lemma Union_mono: "A \<subseteq> B \<Longrightarrow> \<Union>A \<subseteq> \<Union>B"
  1003   by (fact Sup_subset_mono)
  1004 
  1005 
  1006 subsubsection {* Unions of families *}
  1007 
  1008 abbreviation UNION :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
  1009   "UNION \<equiv> SUPR"
  1010 
  1011 text {*
  1012   Note: must use name @{const UNION} here instead of @{text UN}
  1013   to allow the following syntax coexist with the plain constant name.
  1014 *}
  1015 
  1016 syntax
  1017   "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3UN _./ _)" [0, 10] 10)
  1018   "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3UN _:_./ _)" [0, 0, 10] 10)
  1019 
  1020 syntax (xsymbols)
  1021   "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>_./ _)" [0, 10] 10)
  1022   "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>_\<in>_./ _)" [0, 0, 10] 10)
  1023 
  1024 syntax (latex output)
  1025   "_UNION1"     :: "pttrns => 'b set => 'b set"           ("(3\<Union>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
  1026   "_UNION"      :: "pttrn => 'a set => 'b set => 'b set"  ("(3\<Union>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 0, 10] 10)
  1027 
  1028 translations
  1029   "UN x y. B"   == "UN x. UN y. B"
  1030   "UN x. B"     == "CONST UNION CONST UNIV (%x. B)"
  1031   "UN x. B"     == "UN x:CONST UNIV. B"
  1032   "UN x:A. B"   == "CONST UNION A (%x. B)"
  1033 
  1034 text {*
  1035   Note the difference between ordinary xsymbol syntax of indexed
  1036   unions and intersections (e.g.\ @{text"\<Union>a\<^isub>1\<in>A\<^isub>1. B"})
  1037   and their \LaTeX\ rendition: @{term"\<Union>a\<^isub>1\<in>A\<^isub>1. B"}. The
  1038   former does not make the index expression a subscript of the
  1039   union/intersection symbol because this leads to problems with nested
  1040   subscripts in Proof General.
  1041 *}
  1042 
  1043 print_translation {*
  1044   [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}]
  1045 *} -- {* to avoid eta-contraction of body *}
  1046 
  1047 lemma UNION_eq [no_atp]:
  1048   "(\<Union>x\<in>A. B x) = {y. \<exists>x\<in>A. y \<in> B x}"
  1049   by (auto simp add: SUP_def)
  1050 
  1051 lemma bind_UNION [code]:
  1052   "Set.bind A f = UNION A f"
  1053   by (simp add: bind_def UNION_eq)
  1054 
  1055 lemma member_bind [simp]:
  1056   "x \<in> Set.bind P f \<longleftrightarrow> x \<in> UNION P f "
  1057   by (simp add: bind_UNION)
  1058 
  1059 lemma Union_image_eq [simp]:
  1060   "\<Union>(B ` A) = (\<Union>x\<in>A. B x)"
  1061   by (rule sym) (fact SUP_def)
  1062 
  1063 lemma UN_iff [simp]: "b \<in> (\<Union>x\<in>A. B x) \<longleftrightarrow> (\<exists>x\<in>A. b \<in> B x)"
  1064   by (auto simp add: SUP_def image_def)
  1065 
  1066 lemma UN_I [intro]: "a \<in> A \<Longrightarrow> b \<in> B a \<Longrightarrow> b \<in> (\<Union>x\<in>A. B x)"
  1067   -- {* The order of the premises presupposes that @{term A} is rigid;
  1068     @{term b} may be flexible. *}
  1069   by auto
  1070 
  1071 lemma UN_E [elim!]: "b \<in> (\<Union>x\<in>A. B x) \<Longrightarrow> (\<And>x. x\<in>A \<Longrightarrow> b \<in> B x \<Longrightarrow> R) \<Longrightarrow> R"
  1072   by (auto simp add: SUP_def image_def)
  1073 
  1074 lemma UN_cong [cong]:
  1075   "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Union>x\<in>A. C x) = (\<Union>x\<in>B. D x)"
  1076   by (fact SUP_cong)
  1077 
  1078 lemma strong_UN_cong:
  1079   "A = B \<Longrightarrow> (\<And>x. x \<in> B =simp=> C x = D x) \<Longrightarrow> (\<Union>x\<in>A. C x) = (\<Union>x\<in>B. D x)"
  1080   by (unfold simp_implies_def) (fact UN_cong)
  1081 
  1082 lemma image_eq_UN: "f ` A = (\<Union>x\<in>A. {f x})"
  1083   by blast
  1084 
  1085 lemma UN_upper: "a \<in> A \<Longrightarrow> B a \<subseteq> (\<Union>x\<in>A. B x)"
  1086   by (fact SUP_upper)
  1087 
  1088 lemma UN_least: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C) \<Longrightarrow> (\<Union>x\<in>A. B x) \<subseteq> C"
  1089   by (fact SUP_least)
  1090 
  1091 lemma Collect_bex_eq [no_atp]: "{x. \<exists>y\<in>A. P x y} = (\<Union>y\<in>A. {x. P x y})"
  1092   by blast
  1093 
  1094 lemma UN_insert_distrib: "u \<in> A \<Longrightarrow> (\<Union>x\<in>A. insert a (B x)) = insert a (\<Union>x\<in>A. B x)"
  1095   by blast
  1096 
  1097 lemma UN_empty [no_atp]: "(\<Union>x\<in>{}. B x) = {}"
  1098   by (fact SUP_empty)
  1099 
  1100 lemma UN_empty2: "(\<Union>x\<in>A. {}) = {}"
  1101   by (fact SUP_bot) (* already simp *)
  1102 
  1103 lemma UN_absorb: "k \<in> I \<Longrightarrow> A k \<union> (\<Union>i\<in>I. A i) = (\<Union>i\<in>I. A i)"
  1104   by (fact SUP_absorb)
  1105 
  1106 lemma UN_insert [simp]: "(\<Union>x\<in>insert a A. B x) = B a \<union> UNION A B"
  1107   by (fact SUP_insert)
  1108 
  1109 lemma UN_Un [simp]: "(\<Union>i \<in> A \<union> B. M i) = (\<Union>i\<in>A. M i) \<union> (\<Union>i\<in>B. M i)"
  1110   by (fact SUP_union)
  1111 
  1112 lemma UN_UN_flatten: "(\<Union>x \<in> (\<Union>y\<in>A. B y). C x) = (\<Union>y\<in>A. \<Union>x\<in>B y. C x)"
  1113   by blast
  1114 
  1115 lemma UN_subset_iff: "((\<Union>i\<in>I. A i) \<subseteq> B) = (\<forall>i\<in>I. A i \<subseteq> B)"
  1116   by (fact SUP_le_iff)
  1117 
  1118 lemma UN_constant [simp]: "(\<Union>y\<in>A. c) = (if A = {} then {} else c)"
  1119   by (fact SUP_constant)
  1120 
  1121 lemma image_Union: "f ` \<Union>S = (\<Union>x\<in>S. f ` x)"
  1122   by blast
  1123 
  1124 lemma UNION_empty_conv:
  1125   "{} = (\<Union>x\<in>A. B x) \<longleftrightarrow> (\<forall>x\<in>A. B x = {})"
  1126   "(\<Union>x\<in>A. B x) = {} \<longleftrightarrow> (\<forall>x\<in>A. B x = {})"
  1127   by (fact SUP_bot_conv)+ (* already simp *)
  1128 
  1129 lemma Collect_ex_eq [no_atp]: "{x. \<exists>y. P x y} = (\<Union>y. {x. P x y})"
  1130   by blast
  1131 
  1132 lemma ball_UN: "(\<forall>z \<in> UNION A B. P z) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>z \<in> B x. P z)"
  1133   by blast
  1134 
  1135 lemma bex_UN: "(\<exists>z \<in> UNION A B. P z) \<longleftrightarrow> (\<exists>x\<in>A. \<exists>z\<in>B x. P z)"
  1136   by blast
  1137 
  1138 lemma Un_eq_UN: "A \<union> B = (\<Union>b. if b then A else B)"
  1139   by (auto simp add: split_if_mem2)
  1140 
  1141 lemma UN_bool_eq: "(\<Union>b. A b) = (A True \<union> A False)"
  1142   by (fact SUP_UNIV_bool_expand)
  1143 
  1144 lemma UN_Pow_subset: "(\<Union>x\<in>A. Pow (B x)) \<subseteq> Pow (\<Union>x\<in>A. B x)"
  1145   by blast
  1146 
  1147 lemma UN_mono:
  1148   "A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<subseteq> g x) \<Longrightarrow>
  1149     (\<Union>x\<in>A. f x) \<subseteq> (\<Union>x\<in>B. g x)"
  1150   by (fact SUP_subset_mono)
  1151 
  1152 lemma vimage_Union: "f -` (\<Union>A) = (\<Union>X\<in>A. f -` X)"
  1153   by blast
  1154 
  1155 lemma vimage_UN: "f -` (\<Union>x\<in>A. B x) = (\<Union>x\<in>A. f -` B x)"
  1156   by blast
  1157 
  1158 lemma vimage_eq_UN: "f -` B = (\<Union>y\<in>B. f -` {y})"
  1159   -- {* NOT suitable for rewriting *}
  1160   by blast
  1161 
  1162 lemma image_UN: "f ` UNION A B = (\<Union>x\<in>A. f ` B x)"
  1163   by blast
  1164 
  1165 lemma UN_singleton [simp]: "(\<Union>x\<in>A. {x}) = A"
  1166   by blast
  1167 
  1168 
  1169 subsubsection {* Distributive laws *}
  1170 
  1171 lemma Int_Union: "A \<inter> \<Union>B = (\<Union>C\<in>B. A \<inter> C)"
  1172   by (fact inf_Sup)
  1173 
  1174 lemma Un_Inter: "A \<union> \<Inter>B = (\<Inter>C\<in>B. A \<union> C)"
  1175   by (fact sup_Inf)
  1176 
  1177 lemma Int_Union2: "\<Union>B \<inter> A = (\<Union>C\<in>B. C \<inter> A)"
  1178   by (fact Sup_inf)
  1179 
  1180 lemma INT_Int_distrib: "(\<Inter>i\<in>I. A i \<inter> B i) = (\<Inter>i\<in>I. A i) \<inter> (\<Inter>i\<in>I. B i)"
  1181   by (rule sym) (rule INF_inf_distrib)
  1182 
  1183 lemma UN_Un_distrib: "(\<Union>i\<in>I. A i \<union> B i) = (\<Union>i\<in>I. A i) \<union> (\<Union>i\<in>I. B i)"
  1184   by (rule sym) (rule SUP_sup_distrib)
  1185 
  1186 lemma Int_Inter_image: "(\<Inter>x\<in>C. A x \<inter> B x) = \<Inter>(A ` C) \<inter> \<Inter>(B ` C)"
  1187   by (simp only: INT_Int_distrib INF_def)
  1188 
  1189 lemma Un_Union_image: "(\<Union>x\<in>C. A x \<union> B x) = \<Union>(A ` C) \<union> \<Union>(B ` C)"
  1190   -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *}
  1191   -- {* Union of a family of unions *}
  1192   by (simp only: UN_Un_distrib SUP_def)
  1193 
  1194 lemma Un_INT_distrib: "B \<union> (\<Inter>i\<in>I. A i) = (\<Inter>i\<in>I. B \<union> A i)"
  1195   by (fact sup_INF)
  1196 
  1197 lemma Int_UN_distrib: "B \<inter> (\<Union>i\<in>I. A i) = (\<Union>i\<in>I. B \<inter> A i)"
  1198   -- {* Halmos, Naive Set Theory, page 35. *}
  1199   by (fact inf_SUP)
  1200 
  1201 lemma Int_UN_distrib2: "(\<Union>i\<in>I. A i) \<inter> (\<Union>j\<in>J. B j) = (\<Union>i\<in>I. \<Union>j\<in>J. A i \<inter> B j)"
  1202   by (fact SUP_inf_distrib2)
  1203 
  1204 lemma Un_INT_distrib2: "(\<Inter>i\<in>I. A i) \<union> (\<Inter>j\<in>J. B j) = (\<Inter>i\<in>I. \<Inter>j\<in>J. A i \<union> B j)"
  1205   by (fact INF_sup_distrib2)
  1206 
  1207 lemma Union_disjoint: "(\<Union>C \<inter> A = {}) \<longleftrightarrow> (\<forall>B\<in>C. B \<inter> A = {})"
  1208   by (fact Sup_inf_eq_bot_iff)
  1209 
  1210 
  1211 subsubsection {* Complement *}
  1212 
  1213 lemma Compl_INT [simp]: "- (\<Inter>x\<in>A. B x) = (\<Union>x\<in>A. -B x)"
  1214   by (fact uminus_INF)
  1215 
  1216 lemma Compl_UN [simp]: "- (\<Union>x\<in>A. B x) = (\<Inter>x\<in>A. -B x)"
  1217   by (fact uminus_SUP)
  1218 
  1219 
  1220 subsubsection {* Miniscoping and maxiscoping *}
  1221 
  1222 text {* \medskip Miniscoping: pushing in quantifiers and big Unions
  1223            and Intersections. *}
  1224 
  1225 lemma UN_simps [simp]:
  1226   "\<And>a B C. (\<Union>x\<in>C. insert a (B x)) = (if C={} then {} else insert a (\<Union>x\<in>C. B x))"
  1227   "\<And>A B C. (\<Union>x\<in>C. A x \<union> B) = ((if C={} then {} else (\<Union>x\<in>C. A x) \<union> B))"
  1228   "\<And>A B C. (\<Union>x\<in>C. A \<union> B x) = ((if C={} then {} else A \<union> (\<Union>x\<in>C. B x)))"
  1229   "\<And>A B C. (\<Union>x\<in>C. A x \<inter> B) = ((\<Union>x\<in>C. A x) \<inter> B)"
  1230   "\<And>A B C. (\<Union>x\<in>C. A \<inter> B x) = (A \<inter>(\<Union>x\<in>C. B x))"
  1231   "\<And>A B C. (\<Union>x\<in>C. A x - B) = ((\<Union>x\<in>C. A x) - B)"
  1232   "\<And>A B C. (\<Union>x\<in>C. A - B x) = (A - (\<Inter>x\<in>C. B x))"
  1233   "\<And>A B. (\<Union>x\<in>\<Union>A. B x) = (\<Union>y\<in>A. \<Union>x\<in>y. B x)"
  1234   "\<And>A B C. (\<Union>z\<in>UNION A B. C z) = (\<Union>x\<in>A. \<Union>z\<in>B x. C z)"
  1235   "\<And>A B f. (\<Union>x\<in>f`A. B x) = (\<Union>a\<in>A. B (f a))"
  1236   by auto
  1237 
  1238 lemma INT_simps [simp]:
  1239   "\<And>A B C. (\<Inter>x\<in>C. A x \<inter> B) = (if C={} then UNIV else (\<Inter>x\<in>C. A x) \<inter> B)"
  1240   "\<And>A B C. (\<Inter>x\<in>C. A \<inter> B x) = (if C={} then UNIV else A \<inter>(\<Inter>x\<in>C. B x))"
  1241   "\<And>A B C. (\<Inter>x\<in>C. A x - B) = (if C={} then UNIV else (\<Inter>x\<in>C. A x) - B)"
  1242   "\<And>A B C. (\<Inter>x\<in>C. A - B x) = (if C={} then UNIV else A - (\<Union>x\<in>C. B x))"
  1243   "\<And>a B C. (\<Inter>x\<in>C. insert a (B x)) = insert a (\<Inter>x\<in>C. B x)"
  1244   "\<And>A B C. (\<Inter>x\<in>C. A x \<union> B) = ((\<Inter>x\<in>C. A x) \<union> B)"
  1245   "\<And>A B C. (\<Inter>x\<in>C. A \<union> B x) = (A \<union> (\<Inter>x\<in>C. B x))"
  1246   "\<And>A B. (\<Inter>x\<in>\<Union>A. B x) = (\<Inter>y\<in>A. \<Inter>x\<in>y. B x)"
  1247   "\<And>A B C. (\<Inter>z\<in>UNION A B. C z) = (\<Inter>x\<in>A. \<Inter>z\<in>B x. C z)"
  1248   "\<And>A B f. (\<Inter>x\<in>f`A. B x) = (\<Inter>a\<in>A. B (f a))"
  1249   by auto
  1250 
  1251 lemma UN_ball_bex_simps [simp, no_atp]:
  1252   "\<And>A P. (\<forall>x\<in>\<Union>A. P x) \<longleftrightarrow> (\<forall>y\<in>A. \<forall>x\<in>y. P x)"
  1253   "\<And>A B P. (\<forall>x\<in>UNION A B. P x) = (\<forall>a\<in>A. \<forall>x\<in> B a. P x)"
  1254   "\<And>A P. (\<exists>x\<in>\<Union>A. P x) \<longleftrightarrow> (\<exists>y\<in>A. \<exists>x\<in>y. P x)"
  1255   "\<And>A B P. (\<exists>x\<in>UNION A B. P x) \<longleftrightarrow> (\<exists>a\<in>A. \<exists>x\<in>B a. P x)"
  1256   by auto
  1257 
  1258 
  1259 text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *}
  1260 
  1261 lemma UN_extend_simps:
  1262   "\<And>a B C. insert a (\<Union>x\<in>C. B x) = (if C={} then {a} else (\<Union>x\<in>C. insert a (B x)))"
  1263   "\<And>A B C. (\<Union>x\<in>C. A x) \<union> B = (if C={} then B else (\<Union>x\<in>C. A x \<union> B))"
  1264   "\<And>A B C. A \<union> (\<Union>x\<in>C. B x) = (if C={} then A else (\<Union>x\<in>C. A \<union> B x))"
  1265   "\<And>A B C. ((\<Union>x\<in>C. A x) \<inter> B) = (\<Union>x\<in>C. A x \<inter> B)"
  1266   "\<And>A B C. (A \<inter> (\<Union>x\<in>C. B x)) = (\<Union>x\<in>C. A \<inter> B x)"
  1267   "\<And>A B C. ((\<Union>x\<in>C. A x) - B) = (\<Union>x\<in>C. A x - B)"
  1268   "\<And>A B C. (A - (\<Inter>x\<in>C. B x)) = (\<Union>x\<in>C. A - B x)"
  1269   "\<And>A B. (\<Union>y\<in>A. \<Union>x\<in>y. B x) = (\<Union>x\<in>\<Union>A. B x)"
  1270   "\<And>A B C. (\<Union>x\<in>A. \<Union>z\<in>B x. C z) = (\<Union>z\<in>UNION A B. C z)"
  1271   "\<And>A B f. (\<Union>a\<in>A. B (f a)) = (\<Union>x\<in>f`A. B x)"
  1272   by auto
  1273 
  1274 lemma INT_extend_simps:
  1275   "\<And>A B C. (\<Inter>x\<in>C. A x) \<inter> B = (if C={} then B else (\<Inter>x\<in>C. A x \<inter> B))"
  1276   "\<And>A B C. A \<inter> (\<Inter>x\<in>C. B x) = (if C={} then A else (\<Inter>x\<in>C. A \<inter> B x))"
  1277   "\<And>A B C. (\<Inter>x\<in>C. A x) - B = (if C={} then UNIV - B else (\<Inter>x\<in>C. A x - B))"
  1278   "\<And>A B C. A - (\<Union>x\<in>C. B x) = (if C={} then A else (\<Inter>x\<in>C. A - B x))"
  1279   "\<And>a B C. insert a (\<Inter>x\<in>C. B x) = (\<Inter>x\<in>C. insert a (B x))"
  1280   "\<And>A B C. ((\<Inter>x\<in>C. A x) \<union> B) = (\<Inter>x\<in>C. A x \<union> B)"
  1281   "\<And>A B C. A \<union> (\<Inter>x\<in>C. B x) = (\<Inter>x\<in>C. A \<union> B x)"
  1282   "\<And>A B. (\<Inter>y\<in>A. \<Inter>x\<in>y. B x) = (\<Inter>x\<in>\<Union>A. B x)"
  1283   "\<And>A B C. (\<Inter>x\<in>A. \<Inter>z\<in>B x. C z) = (\<Inter>z\<in>UNION A B. C z)"
  1284   "\<And>A B f. (\<Inter>a\<in>A. B (f a)) = (\<Inter>x\<in>f`A. B x)"
  1285   by auto
  1286 
  1287 text {* Finally *}
  1288 
  1289 no_notation
  1290   less_eq (infix "\<sqsubseteq>" 50) and
  1291   less (infix "\<sqsubset>" 50)
  1292 
  1293 lemmas mem_simps =
  1294   insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff
  1295   mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff
  1296   -- {* Each of these has ALREADY been added @{text "[simp]"} above. *}
  1297 
  1298 end
  1299