--- a/src/HOL/Set.thy Thu Mar 05 08:23:10 2009 +0100
+++ b/src/HOL/Set.thy Thu Mar 05 08:23:11 2009 +0100
@@ -1,12 +1,11 @@
(* Title: HOL/Set.thy
- ID: $Id$
Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
*)
header {* Set theory for higher-order logic *}
theory Set
-imports Orderings
+imports Lattices
begin
text {* A set in HOL is simply a predicate. *}
@@ -19,36 +18,21 @@
types 'a set = "'a => bool"
consts
- "{}" :: "'a set" ("{}")
- UNIV :: "'a set"
+ Collect :: "('a => bool) => 'a set" -- "comprehension"
+ "op :" :: "'a => 'a set => bool" -- "membership"
insert :: "'a => 'a set => 'a set"
- Collect :: "('a => bool) => 'a set" -- "comprehension"
- "op Int" :: "'a set => 'a set => 'a set" (infixl "Int" 70)
- "op Un" :: "'a set => 'a set => 'a set" (infixl "Un" 65)
- UNION :: "'a set => ('a => 'b set) => 'b set" -- "general union"
- INTER :: "'a set => ('a => 'b set) => 'b set" -- "general intersection"
- Union :: "'a set set => 'a set" -- "union of a set"
- Inter :: "'a set set => 'a set" -- "intersection of a set"
- Pow :: "'a set => 'a set set" -- "powerset"
Ball :: "'a set => ('a => bool) => bool" -- "bounded universal quantifiers"
Bex :: "'a set => ('a => bool) => bool" -- "bounded existential quantifiers"
Bex1 :: "'a set => ('a => bool) => bool" -- "bounded unique existential quantifiers"
+ Pow :: "'a set => 'a set set" -- "powerset"
image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90)
- "op :" :: "'a => 'a set => bool" -- "membership"
+
+local
notation
"op :" ("op :") and
"op :" ("(_/ : _)" [50, 51] 50)
-local
-
-
-subsection {* Additional concrete syntax *}
-
-abbreviation
- range :: "('a => 'b) => 'b set" where -- "of function"
- "range f == f ` UNIV"
-
abbreviation
"not_mem x A == ~ (x : A)" -- "non-membership"
@@ -57,32 +41,51 @@
not_mem ("(_/ ~: _)" [50, 51] 50)
notation (xsymbols)
- "op Int" (infixl "\<inter>" 70) and
- "op Un" (infixl "\<union>" 65) and
"op :" ("op \<in>") and
"op :" ("(_/ \<in> _)" [50, 51] 50) and
not_mem ("op \<notin>") and
- not_mem ("(_/ \<notin> _)" [50, 51] 50) and
- Union ("\<Union>_" [90] 90) and
- Inter ("\<Inter>_" [90] 90)
+ not_mem ("(_/ \<notin> _)" [50, 51] 50)
notation (HTML output)
- "op Int" (infixl "\<inter>" 70) and
- "op Un" (infixl "\<union>" 65) and
"op :" ("op \<in>") and
"op :" ("(_/ \<in> _)" [50, 51] 50) and
not_mem ("op \<notin>") and
not_mem ("(_/ \<notin> _)" [50, 51] 50)
syntax
+ "@Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})")
+
+translations
+ "{x. P}" == "Collect (%x. P)"
+
+definition empty :: "'a set" ("{}") where
+ "empty \<equiv> {x. False}"
+
+definition UNIV :: "'a set" where
+ "UNIV \<equiv> {x. True}"
+
+syntax
"@Finset" :: "args => 'a set" ("{(_)}")
- "@Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})")
- "@SetCompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})")
- "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ :/ _./ _})")
- "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10)
- "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10)
- "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 10] 10)
- "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10)
+
+translations
+ "{x, xs}" == "insert x {xs}"
+ "{x}" == "insert x {}"
+
+definition Int :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Int" 70) where
+ "A Int B \<equiv> {x. x \<in> A \<and> x \<in> B}"
+
+definition Un :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Un" 65) where
+ "A Un B \<equiv> {x. x \<in> A \<or> x \<in> B}"
+
+notation (xsymbols)
+ "Int" (infixl "\<inter>" 70) and
+ "Un" (infixl "\<union>" 65)
+
+notation (HTML output)
+ "Int" (infixl "\<inter>" 70) and
+ "Un" (infixl "\<union>" 65)
+
+syntax
"_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10)
"_Bex" :: "pttrn => 'a set => bool => bool" ("(3EX _:_./ _)" [0, 0, 10] 10)
"_Bex1" :: "pttrn => 'a set => bool => bool" ("(3EX! _:_./ _)" [0, 0, 10] 10)
@@ -93,24 +96,6 @@
"_Bex" :: "pttrn => 'a set => bool => bool" ("(3? _:_./ _)" [0, 0, 10] 10)
"_Bex1" :: "pttrn => 'a set => bool => bool" ("(3?! _:_./ _)" [0, 0, 10] 10)
-translations
- "{x, xs}" == "insert x {xs}"
- "{x}" == "insert x {}"
- "{x. P}" == "Collect (%x. P)"
- "{x:A. P}" => "{x. x:A & P}"
- "UN x y. B" == "UN x. UN y. B"
- "UN x. B" == "UNION UNIV (%x. B)"
- "UN x. B" == "UN x:UNIV. B"
- "INT x y. B" == "INT x. INT y. B"
- "INT x. B" == "INTER UNIV (%x. B)"
- "INT x. B" == "INT x:UNIV. B"
- "UN x:A. B" == "UNION A (%x. B)"
- "INT x:A. B" == "INTER A (%x. B)"
- "ALL x:A. P" == "Ball A (%x. P)"
- "EX x:A. P" == "Bex A (%x. P)"
- "EX! x:A. P" == "Bex1 A (%x. P)"
- "LEAST x:A. P" => "LEAST x. x:A & P"
-
syntax (xsymbols)
"_Ball" :: "pttrn => 'a set => bool => bool" ("(3\<forall>_\<in>_./ _)" [0, 0, 10] 10)
"_Bex" :: "pttrn => 'a set => bool => bool" ("(3\<exists>_\<in>_./ _)" [0, 0, 10] 10)
@@ -122,26 +107,71 @@
"_Bex" :: "pttrn => 'a set => bool => bool" ("(3\<exists>_\<in>_./ _)" [0, 0, 10] 10)
"_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\<exists>!_\<in>_./ _)" [0, 0, 10] 10)
+translations
+ "ALL x:A. P" == "Ball A (%x. P)"
+ "EX x:A. P" == "Bex A (%x. P)"
+ "EX! x:A. P" == "Bex1 A (%x. P)"
+ "LEAST x:A. P" => "LEAST x. x:A & P"
+
+definition INTER :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
+ "INTER A B \<equiv> {y. \<forall>x\<in>A. y \<in> B x}"
+
+definition UNION :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b set) \<Rightarrow> 'b set" where
+ "UNION A B \<equiv> {y. \<exists>x\<in>A. y \<in> B x}"
+
+definition Inter :: "'a set set \<Rightarrow> 'a set" where
+ "Inter S \<equiv> INTER S (\<lambda>x. x)"
+
+definition Union :: "'a set set \<Rightarrow> 'a set" where
+ "Union S \<equiv> UNION S (\<lambda>x. x)"
+
+notation (xsymbols)
+ Inter ("\<Inter>_" [90] 90) and
+ Union ("\<Union>_" [90] 90)
+
+
+subsection {* Additional concrete syntax *}
+
+syntax
+ "@SetCompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})")
+ "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ :/ _./ _})")
+ "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10)
+ "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10)
+ "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 10] 10)
+ "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10)
+
syntax (xsymbols)
"@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ \<in>/ _./ _})")
+ "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\<Inter>_./ _)" [0, 10] 10)
"@UNION1" :: "pttrns => 'b set => 'b set" ("(3\<Union>_./ _)" [0, 10] 10)
- "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\<Inter>_./ _)" [0, 10] 10)
+ "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\<Inter>_\<in>_./ _)" [0, 10] 10)
"@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\<Union>_\<in>_./ _)" [0, 10] 10)
- "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\<Inter>_\<in>_./ _)" [0, 10] 10)
syntax (latex output)
+ "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\<Inter>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
"@UNION1" :: "pttrns => 'b set => 'b set" ("(3\<Union>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
- "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\<Inter>(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10)
+ "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 10] 10)
"@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\<Union>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 10] 10)
- "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\<Inter>(00\<^bsub>_\<in>_\<^esub>)/ _)" [0, 10] 10)
-
-text{*
+
+translations
+ "{x:A. P}" => "{x. x:A & P}"
+ "INT x y. B" == "INT x. INT y. B"
+ "INT x. B" == "CONST INTER CONST UNIV (%x. B)"
+ "INT x. B" == "INT x:CONST UNIV. B"
+ "INT x:A. B" == "CONST INTER A (%x. B)"
+ "UN x y. B" == "UN x. UN y. B"
+ "UN x. B" == "CONST UNION CONST UNIV (%x. B)"
+ "UN x. B" == "UN x:CONST UNIV. B"
+ "UN x:A. B" == "CONST UNION A (%x. B)"
+
+text {*
Note the difference between ordinary xsymbol syntax of indexed
unions and intersections (e.g.\ @{text"\<Union>a\<^isub>1\<in>A\<^isub>1. B"})
and their \LaTeX\ rendition: @{term"\<Union>a\<^isub>1\<in>A\<^isub>1. B"}. The
former does not make the index expression a subscript of the
union/intersection symbol because this leads to problems with nested
- subscripts in Proof General. *}
+ subscripts in Proof General.
+*}
abbreviation
subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
@@ -183,6 +213,10 @@
supset_eq ("op \<supseteq>") and
supset_eq ("(_/ \<supseteq> _)" [50, 51] 50)
+abbreviation
+ range :: "('a => 'b) => 'b set" where -- "of function"
+ "range f == f ` UNIV"
+
subsubsection "Bounded quantifiers"
@@ -280,12 +314,12 @@
(* To avoid eta-contraction of body: *)
print_translation {*
let
- fun btr' syn [A,Abs abs] =
- let val (x,t) = atomic_abs_tr' abs
+ fun btr' syn [A, Abs abs] =
+ let val (x, t) = atomic_abs_tr' abs
in Syntax.const syn $ x $ A $ t end
in
-[("Ball", btr' "_Ball"),("Bex", btr' "_Bex"),
- ("UNION", btr' "@UNION"),("INTER", btr' "@INTER")]
+[(@{const_syntax Ball}, btr' "_Ball"), (@{const_syntax Bex}, btr' "_Bex"),
+ (@{const_syntax UNION}, btr' "@UNION"),(@{const_syntax INTER}, btr' "@INTER")]
end
*}
@@ -373,15 +407,7 @@
end
defs
- Un_def: "A Un B == {x. x:A | x:B}"
- Int_def: "A Int B == {x. x:A & x:B}"
- INTER_def: "INTER A B == {y. ALL x:A. y: B(x)}"
- UNION_def: "UNION A B == {y. EX x:A. y: B(x)}"
- Inter_def: "Inter S == (INT x:S. x)"
- Union_def: "Union S == (UN x:S. x)"
Pow_def: "Pow A == {B. B <= A}"
- empty_def: "{} == {x. False}"
- UNIV_def: "UNIV == {x. True}"
insert_def: "insert a B == {x. x=a} Un B"
image_def: "f`A == {y. EX x:A. y = f(x)}"
@@ -1048,12 +1074,12 @@
to use term-nets to associate patterns with rules. Also, if a rule fails to
apply, then the formula should be kept.
[("HOL.uminus", Compl_iff RS iffD1), ("HOL.minus", [Diff_iff RS iffD1]),
- ("op Int", [IntD1,IntD2]),
+ ("Int", [IntD1,IntD2]),
("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])]
*)
ML {*
- val mksimps_pairs = [("Ball", @{thms bspec})] @ mksimps_pairs;
+ val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs;
*}
declaration {* fn _ =>
Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs))
@@ -2160,9 +2186,7 @@
subsection {* Getting the Contents of a Singleton Set *}
-definition
- contents :: "'a set \<Rightarrow> 'a"
-where
+definition contents :: "'a set \<Rightarrow> 'a" where
[code del]: "contents X = (THE x. X = {x})"
lemma contents_eq [simp]: "contents {x} = x"
@@ -2215,6 +2239,255 @@
unfolding vimage_def Collect_def mem_def ..
+subsection {* Complete lattices *}
+
+notation
+ less_eq (infix "\<sqsubseteq>" 50) and
+ less (infix "\<sqsubset>" 50) and
+ inf (infixl "\<sqinter>" 70) and
+ sup (infixl "\<squnion>" 65)
+
+class complete_lattice = lattice + bot + top +
+ fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
+ and Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
+ assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x"
+ and Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A"
+ assumes Sup_upper: "x \<in> A \<Longrightarrow> x \<sqsubseteq> \<Squnion>A"
+ and Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> z) \<Longrightarrow> \<Squnion>A \<sqsubseteq> z"
+begin
+
+lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<le> a}"
+ by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
+
+lemma Sup_Inf: "\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<le> b}"
+ by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
+
+lemma Inf_Univ: "\<Sqinter>UNIV = \<Squnion>{}"
+ unfolding Sup_Inf by auto
+
+lemma Sup_Univ: "\<Squnion>UNIV = \<Sqinter>{}"
+ unfolding Inf_Sup by auto
+
+lemma Inf_insert: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
+ by (auto intro: antisym Inf_greatest Inf_lower)
+
+lemma Sup_insert: "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
+ by (auto intro: antisym Sup_least Sup_upper)
+
+lemma Inf_singleton [simp]:
+ "\<Sqinter>{a} = a"
+ by (auto intro: antisym Inf_lower Inf_greatest)
+
+lemma Sup_singleton [simp]:
+ "\<Squnion>{a} = a"
+ by (auto intro: antisym Sup_upper Sup_least)
+
+lemma Inf_insert_simp:
+ "\<Sqinter>insert a A = (if A = {} then a else a \<sqinter> \<Sqinter>A)"
+ by (cases "A = {}") (simp_all, simp add: Inf_insert)
+
+lemma Sup_insert_simp:
+ "\<Squnion>insert a A = (if A = {} then a else a \<squnion> \<Squnion>A)"
+ by (cases "A = {}") (simp_all, simp add: Sup_insert)
+
+lemma Inf_binary:
+ "\<Sqinter>{a, b} = a \<sqinter> b"
+ by (simp add: Inf_insert_simp)
+
+lemma Sup_binary:
+ "\<Squnion>{a, b} = a \<squnion> b"
+ by (simp add: Sup_insert_simp)
+
+lemma bot_def:
+ "bot = \<Squnion>{}"
+ by (auto intro: antisym Sup_least)
+
+lemma top_def:
+ "top = \<Sqinter>{}"
+ by (auto intro: antisym Inf_greatest)
+
+lemma sup_bot [simp]:
+ "x \<squnion> bot = x"
+ using bot_least [of x] by (simp add: le_iff_sup sup_commute)
+
+lemma inf_top [simp]:
+ "x \<sqinter> top = x"
+ using top_greatest [of x] by (simp add: le_iff_inf inf_commute)
+
+definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
+ "SUPR A f == \<Squnion> (f ` A)"
+
+definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
+ "INFI A f == \<Sqinter> (f ` A)"
+
+end
+
+syntax
+ "_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10)
+ "_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10)
+ "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10)
+ "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10)
+
+translations
+ "SUP x y. B" == "SUP x. SUP y. B"
+ "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)"
+ "SUP x. B" == "SUP x:CONST UNIV. B"
+ "SUP x:A. B" == "CONST SUPR A (%x. B)"
+ "INF x y. B" == "INF x. INF y. B"
+ "INF x. B" == "CONST INFI CONST UNIV (%x. B)"
+ "INF x. B" == "INF x:CONST UNIV. B"
+ "INF x:A. B" == "CONST INFI A (%x. B)"
+
+(* To avoid eta-contraction of body: *)
+print_translation {*
+let
+ fun btr' syn (A :: Abs abs :: ts) =
+ let val (x,t) = atomic_abs_tr' abs
+ in list_comb (Syntax.const syn $ x $ A $ t, ts) end
+ val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const
+in
+[(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")]
+end
+*}
+
+context complete_lattice
+begin
+
+lemma le_SUPI: "i : A \<Longrightarrow> M i \<le> (SUP i:A. M i)"
+ by (auto simp add: SUPR_def intro: Sup_upper)
+
+lemma SUP_leI: "(\<And>i. i : A \<Longrightarrow> M i \<le> u) \<Longrightarrow> (SUP i:A. M i) \<le> u"
+ by (auto simp add: SUPR_def intro: Sup_least)
+
+lemma INF_leI: "i : A \<Longrightarrow> (INF i:A. M i) \<le> M i"
+ by (auto simp add: INFI_def intro: Inf_lower)
+
+lemma le_INFI: "(\<And>i. i : A \<Longrightarrow> u \<le> M i) \<Longrightarrow> u \<le> (INF i:A. M i)"
+ by (auto simp add: INFI_def intro: Inf_greatest)
+
+lemma SUP_const[simp]: "A \<noteq> {} \<Longrightarrow> (SUP i:A. M) = M"
+ by (auto intro: antisym SUP_leI le_SUPI)
+
+lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M"
+ by (auto intro: antisym INF_leI le_INFI)
+
+end
+
+
+subsection {* Bool as complete lattice *}
+
+instantiation bool :: complete_lattice
+begin
+
+definition
+ Inf_bool_def: "\<Sqinter>A \<longleftrightarrow> (\<forall>x\<in>A. x)"
+
+definition
+ Sup_bool_def: "\<Squnion>A \<longleftrightarrow> (\<exists>x\<in>A. x)"
+
+instance
+ by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def)
+
+end
+
+lemma Inf_empty_bool [simp]:
+ "\<Sqinter>{}"
+ unfolding Inf_bool_def by auto
+
+lemma not_Sup_empty_bool [simp]:
+ "\<not> Sup {}"
+ unfolding Sup_bool_def by auto
+
+
+subsection {* Fun as complete lattice *}
+
+instantiation "fun" :: (type, complete_lattice) complete_lattice
+begin
+
+definition
+ Inf_fun_def [code del]: "\<Sqinter>A = (\<lambda>x. \<Sqinter>{y. \<exists>f\<in>A. y = f x})"
+
+definition
+ Sup_fun_def [code del]: "\<Squnion>A = (\<lambda>x. \<Squnion>{y. \<exists>f\<in>A. y = f x})"
+
+instance
+ by intro_classes
+ (auto simp add: Inf_fun_def Sup_fun_def le_fun_def
+ intro: Inf_lower Sup_upper Inf_greatest Sup_least)
+
+end
+
+lemma Inf_empty_fun:
+ "\<Sqinter>{} = (\<lambda>_. \<Sqinter>{})"
+ by rule (auto simp add: Inf_fun_def)
+
+lemma Sup_empty_fun:
+ "\<Squnion>{} = (\<lambda>_. \<Squnion>{})"
+ by rule (auto simp add: Sup_fun_def)
+
+
+subsection {* Set as lattice *}
+
+lemma inf_set_eq: "A \<sqinter> B = A \<inter> B"
+ apply (rule subset_antisym)
+ apply (rule Int_greatest)
+ apply (rule inf_le1)
+ apply (rule inf_le2)
+ apply (rule inf_greatest)
+ apply (rule Int_lower1)
+ apply (rule Int_lower2)
+ done
+
+lemma sup_set_eq: "A \<squnion> B = A \<union> B"
+ apply (rule subset_antisym)
+ apply (rule sup_least)
+ apply (rule Un_upper1)
+ apply (rule Un_upper2)
+ apply (rule Un_least)
+ apply (rule sup_ge1)
+ apply (rule sup_ge2)
+ done
+
+lemma mono_Int: "mono f \<Longrightarrow> f (A \<inter> B) \<subseteq> f A \<inter> f B"
+ apply (fold inf_set_eq sup_set_eq)
+ apply (erule mono_inf)
+ done
+
+lemma mono_Un: "mono f \<Longrightarrow> f A \<union> f B \<subseteq> f (A \<union> B)"
+ apply (fold inf_set_eq sup_set_eq)
+ apply (erule mono_sup)
+ done
+
+lemma Inf_set_eq: "\<Sqinter>S = \<Inter>S"
+ apply (rule subset_antisym)
+ apply (rule Inter_greatest)
+ apply (erule Inf_lower)
+ apply (rule Inf_greatest)
+ apply (erule Inter_lower)
+ done
+
+lemma Sup_set_eq: "\<Squnion>S = \<Union>S"
+ apply (rule subset_antisym)
+ apply (rule Sup_least)
+ apply (erule Union_upper)
+ apply (rule Union_least)
+ apply (erule Sup_upper)
+ done
+
+lemma top_set_eq: "top = UNIV"
+ by (iprover intro!: subset_antisym subset_UNIV top_greatest)
+
+lemma bot_set_eq: "bot = {}"
+ by (iprover intro!: subset_antisym empty_subsetI bot_least)
+
+no_notation
+ less_eq (infix "\<sqsubseteq>" 50) and
+ less (infix "\<sqsubset>" 50) and
+ inf (infixl "\<sqinter>" 70) and
+ sup (infixl "\<squnion>" 65) and
+ Inf ("\<Sqinter>_" [900] 900) and
+ Sup ("\<Squnion>_" [900] 900)
+
subsection {* Basic ML bindings *}