simplified HOL bootstrap
authorhaftmann
Fri, 20 Jul 2007 14:27:56 +0200
changeset 23878 bd651ecd4b8a
parent 23877 307f75aaefca
child 23879 4776af8be741
simplified HOL bootstrap
src/HOL/ATP_Linkup.thy
src/HOL/Finite_Set.thy
src/HOL/FixedPoint.thy
src/HOL/Fun.thy
src/HOL/HOL.thy
src/HOL/Lattices.thy
src/HOL/Set.thy
--- a/src/HOL/ATP_Linkup.thy	Fri Jul 20 00:01:40 2007 +0200
+++ b/src/HOL/ATP_Linkup.thy	Fri Jul 20 14:27:56 2007 +0200
@@ -7,7 +7,7 @@
 header{* The Isabelle-ATP Linkup *}
 
 theory ATP_Linkup
-imports Map Hilbert_Choice
+imports Divides Hilbert_Choice Record
 uses
   "Tools/polyhash.ML"
   "Tools/res_clause.ML"
--- a/src/HOL/Finite_Set.thy	Fri Jul 20 00:01:40 2007 +0200
+++ b/src/HOL/Finite_Set.thy	Fri Jul 20 14:27:56 2007 +0200
@@ -7,7 +7,7 @@
 header {* Finite sets *}
 
 theory Finite_Set
-imports Divides Equiv_Relations IntDef
+imports IntDef Divides
 begin
 
 subsection {* Definition and basic properties *}
@@ -94,6 +94,7 @@
   qed
 qed
 
+
 text{* Finite sets are the images of initial segments of natural numbers: *}
 
 lemma finite_imp_nat_seg_image_inj_on:
--- a/src/HOL/FixedPoint.thy	Fri Jul 20 00:01:40 2007 +0200
+++ b/src/HOL/FixedPoint.thy	Fri Jul 20 14:27:56 2007 +0200
@@ -8,298 +8,9 @@
 header {* Fixed Points and the Knaster-Tarski Theorem*}
 
 theory FixedPoint
-imports Fun
-begin
-
-subsection {* Complete lattices *}
-
-class complete_lattice = lattice +
-  fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
-  assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x"
-  assumes Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A"
+imports Lattices
 begin
 
-definition
-  Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
-where
-  "\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<^loc>\<le> b}"
-
-lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<^loc>\<le> a}"
-  unfolding Sup_def by (auto intro: Inf_greatest Inf_lower)
-
-lemma Sup_upper: "x \<in> A \<Longrightarrow> x \<^loc>\<le> \<Squnion>A"
-  by (auto simp: Sup_def intro: Inf_greatest)
-
-lemma Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<^loc>\<le> z) \<Longrightarrow> \<Squnion>A \<^loc>\<le> z"
-  by (auto simp: Sup_def intro: Inf_lower)
-
-lemma Inf_Univ: "\<Sqinter>UNIV = \<Squnion>{}"
-  unfolding Sup_def by auto
-
-lemma Sup_Univ: "\<Squnion>UNIV = \<Sqinter>{}"
-  unfolding Inf_Sup by auto
-
-lemma Inf_insert: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
-  apply (rule antisym)
-  apply (rule le_infI)
-  apply (rule Inf_lower)
-  apply simp
-  apply (rule Inf_greatest)
-  apply (rule Inf_lower)
-  apply simp
-  apply (rule Inf_greatest)
-  apply (erule insertE)
-  apply (rule le_infI1)
-  apply simp
-  apply (rule le_infI2)
-  apply (erule Inf_lower)
-  done
-
-lemma Sup_insert: "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
-  apply (rule antisym)
-  apply (rule Sup_least)
-  apply (erule insertE)
-  apply (rule le_supI1)
-  apply simp
-  apply (rule le_supI2)
-  apply (erule Sup_upper)
-  apply (rule le_supI)
-  apply (rule Sup_upper)
-  apply simp
-  apply (rule Sup_least)
-  apply (rule Sup_upper)
-  apply simp
-  done
-
-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)
-
-end
-
-lemmas Sup_def = Sup_def [folded complete_lattice_class.Sup]
-lemmas Sup_upper = Sup_upper [folded complete_lattice_class.Sup]
-lemmas Sup_least = Sup_least [folded complete_lattice_class.Sup]
-
-lemmas Sup_insert [code func] = Sup_insert [folded complete_lattice_class.Sup]
-lemmas Sup_singleton [simp, code func] = Sup_singleton [folded complete_lattice_class.Sup]
-lemmas Sup_insert_simp = Sup_insert_simp [folded complete_lattice_class.Sup]
-lemmas Sup_binary = Sup_binary [folded complete_lattice_class.Sup]
-
-(* FIXME: definition inside class does not work *)
-definition
-  top :: "'a::complete_lattice"
-where
-  "top = Inf {}"
-
-definition
-  bot :: "'a::complete_lattice"
-where
-  "bot = Sup {}"
-
-lemma top_greatest [simp]: "x \<le> top"
-  by (unfold top_def, rule Inf_greatest, simp)
-
-lemma bot_least [simp]: "bot \<le> x"
-  by (unfold bot_def, rule Sup_least, simp)
-
-definition
-  SUPR :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b" where
-  "SUPR A f == Sup (f ` A)"
-
-definition
-  INFI :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b" where
-  "INFI A f == Inf (f ` A)"
-
-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 UNIV (%x. B)"
-  "SUP x. B"     == "SUP x: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 UNIV (%x. B)"
-  "INF x. B"     == "INF x: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
-*}
-
-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 mono_inf: "mono f \<Longrightarrow> f (inf A B) <= inf (f A) (f B)"
-  by (auto simp add: mono_def)
-
-lemma mono_sup: "mono f \<Longrightarrow> sup (f A) (f B) <= f (sup A B)"
-  by (auto simp add: mono_def)
-
-lemma SUP_const[simp]: "A \<noteq> {} \<Longrightarrow> (SUP i:A. M) = M"
-  by (auto intro: order_antisym SUP_leI le_SUPI)
-
-lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M"
-  by (auto intro: order_antisym INF_leI le_INFI)
-
-
-subsection {* Some instances of the type class of complete lattices *}
-
-subsubsection {* Booleans *}
-
-instance bool :: complete_lattice
-  Inf_bool_def: "Inf A \<equiv> \<forall>x\<in>A. x"
-  apply intro_classes
-  apply (unfold Inf_bool_def)
-  apply (iprover intro!: le_boolI elim: ballE)
-  apply (iprover intro!: ballI le_boolI elim: ballE le_boolE)
-  done
-
-theorem Sup_bool_eq: "Sup A \<longleftrightarrow> (\<exists>x\<in>A. x)"
-  apply (rule order_antisym)
-  apply (rule Sup_least)
-  apply (rule le_boolI)
-  apply (erule bexI, assumption)
-  apply (rule le_boolI)
-  apply (erule bexE)
-  apply (rule le_boolE)
-  apply (rule Sup_upper)
-  apply assumption+
-  done
-
-lemma Inf_empty_bool [simp]:
-  "Inf {}"
-  unfolding Inf_bool_def by auto
-
-lemma not_Sup_empty_bool [simp]:
-  "\<not> Sup {}"
-  unfolding Sup_def Inf_bool_def by auto
-
-lemma top_bool_eq: "top = True"
-  by (iprover intro!: order_antisym le_boolI top_greatest)
-
-lemma bot_bool_eq: "bot = False"
-  by (iprover intro!: order_antisym le_boolI bot_least)
-
-
-subsubsection {* Functions *}
-
-instance "fun" :: (type, complete_lattice) complete_lattice
-  Inf_fun_def: "Inf A \<equiv> (\<lambda>x. Inf {y. \<exists>f\<in>A. y = f x})"
-  apply intro_classes
-  apply (unfold Inf_fun_def)
-  apply (rule le_funI)
-  apply (rule Inf_lower)
-  apply (rule CollectI)
-  apply (rule bexI)
-  apply (rule refl)
-  apply assumption
-  apply (rule le_funI)
-  apply (rule Inf_greatest)
-  apply (erule CollectE)
-  apply (erule bexE)
-  apply (iprover elim: le_funE)
-  done
-
-lemmas [code func del] = Inf_fun_def
-
-theorem Sup_fun_eq: "Sup A = (\<lambda>x. Sup {y. \<exists>f\<in>A. y = f x})"
-  apply (rule order_antisym)
-  apply (rule Sup_least)
-  apply (rule le_funI)
-  apply (rule Sup_upper)
-  apply fast
-  apply (rule le_funI)
-  apply (rule Sup_least)
-  apply (erule CollectE)
-  apply (erule bexE)
-  apply (drule le_funD [OF Sup_upper])
-  apply simp
-  done
-
-lemma Inf_empty_fun:
-  "Inf {} = (\<lambda>_. Inf {})"
-  by rule (auto simp add: Inf_fun_def)
-
-lemma Sup_empty_fun:
-  "Sup {} = (\<lambda>_. Sup {})"
-proof -
-  have aux: "\<And>x. {y. \<exists>f. y = f x} = UNIV" by auto
-  show ?thesis
-  by (auto simp add: Sup_def Inf_fun_def Inf_binary inf_bool_eq aux)
-qed
-
-lemma top_fun_eq: "top = (\<lambda>x. top)"
-  by (iprover intro!: order_antisym le_funI top_greatest)
-
-lemma bot_fun_eq: "bot = (\<lambda>x. bot)"
-  by (iprover intro!: order_antisym le_funI bot_least)
-
-
-subsubsection {* Sets *}
-
-instance set :: (type) complete_lattice
-  Inf_set_def: "Inf S \<equiv> \<Inter>S"
-  by intro_classes (auto simp add: Inf_set_def)
-
-lemmas [code func del] = Inf_set_def
-
-theorem Sup_set_eq: "Sup 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)
-
-
 subsection {* Least and greatest fixed points *}
 
 definition
--- a/src/HOL/Fun.thy	Fri Jul 20 00:01:40 2007 +0200
+++ b/src/HOL/Fun.thy	Fri Jul 20 14:27:56 2007 +0200
@@ -460,87 +460,6 @@
 by (simp add: bij_def)
 
 
-subsection {* Order and lattice on functions *}
-
-instance "fun" :: (type, ord) ord
-  le_fun_def: "f \<le> g \<equiv> \<forall>x. f x \<le> g x"
-  less_fun_def: "f < g \<equiv> f \<le> g \<and> f \<noteq> g" ..
-
-lemmas [code func del] = le_fun_def less_fun_def
-
-instance "fun" :: (type, order) order
-  by default
-    (auto simp add: le_fun_def less_fun_def expand_fun_eq
-       intro: order_trans order_antisym)
-
-lemma le_funI: "(\<And>x. f x \<le> g x) \<Longrightarrow> f \<le> g"
-  unfolding le_fun_def by simp
-
-lemma le_funE: "f \<le> g \<Longrightarrow> (f x \<le> g x \<Longrightarrow> P) \<Longrightarrow> P"
-  unfolding le_fun_def by simp
-
-lemma le_funD: "f \<le> g \<Longrightarrow> f x \<le> g x"
-  unfolding le_fun_def by simp
-
-text {*
-  Handy introduction and elimination rules for @{text "\<le>"}
-  on unary and binary predicates
-*}
-
-lemma predicate1I [Pure.intro!, intro!]:
-  assumes PQ: "\<And>x. P x \<Longrightarrow> Q x"
-  shows "P \<le> Q"
-  apply (rule le_funI)
-  apply (rule le_boolI)
-  apply (rule PQ)
-  apply assumption
-  done
-
-lemma predicate1D [Pure.dest, dest]: "P \<le> Q \<Longrightarrow> P x \<Longrightarrow> Q x"
-  apply (erule le_funE)
-  apply (erule le_boolE)
-  apply assumption+
-  done
-
-lemma predicate2I [Pure.intro!, intro!]:
-  assumes PQ: "\<And>x y. P x y \<Longrightarrow> Q x y"
-  shows "P \<le> Q"
-  apply (rule le_funI)+
-  apply (rule le_boolI)
-  apply (rule PQ)
-  apply assumption
-  done
-
-lemma predicate2D [Pure.dest, dest]: "P \<le> Q \<Longrightarrow> P x y \<Longrightarrow> Q x y"
-  apply (erule le_funE)+
-  apply (erule le_boolE)
-  apply assumption+
-  done
-
-lemma rev_predicate1D: "P x ==> P <= Q ==> Q x"
-  by (rule predicate1D)
-
-lemma rev_predicate2D: "P x y ==> P <= Q ==> Q x y"
-  by (rule predicate2D)
-
-instance "fun" :: (type, lattice) lattice
-  inf_fun_eq: "inf f g \<equiv> (\<lambda>x. inf (f x) (g x))"
-  sup_fun_eq: "sup f g \<equiv> (\<lambda>x. sup (f x) (g x))"
-apply intro_classes
-unfolding inf_fun_eq sup_fun_eq
-apply (auto intro: le_funI)
-apply (rule le_funI)
-apply (auto dest: le_funD)
-apply (rule le_funI)
-apply (auto dest: le_funD)
-done
-
-lemmas [code func del] = inf_fun_eq sup_fun_eq
-
-instance "fun" :: (type, distrib_lattice) distrib_lattice
-  by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1)
-
-
 subsection {* Proof tool setup *} 
 
 text {* simplifies terms of the form
@@ -600,8 +519,6 @@
 val datatype_injI = @{thm datatype_injI}
 val range_ex1_eq = @{thm range_ex1_eq}
 val expand_fun_eq = @{thm expand_fun_eq}
-val sup_fun_eq = @{thm sup_fun_eq}
-val sup_bool_eq = @{thm sup_bool_eq}
 *}
 
 end
--- a/src/HOL/HOL.thy	Fri Jul 20 00:01:40 2007 +0200
+++ b/src/HOL/HOL.thy	Fri Jul 20 14:27:56 2007 +0200
@@ -218,7 +218,6 @@
 class minus = type +
   fixes uminus :: "'a \<Rightarrow> 'a" 
     and minus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "\<^loc>-" 65)
-    and abs :: "'a \<Rightarrow> 'a"
 
 class times = type +
   fixes times :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "\<^loc>*" 70)
@@ -227,6 +226,9 @@
   fixes inverse :: "'a \<Rightarrow> 'a"
     and divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "\<^loc>'/" 70)
 
+class abs = type +
+  fixes abs :: "'a \<Rightarrow> 'a"
+
 notation
   uminus  ("- _" [81] 80)
 
@@ -235,6 +237,70 @@
 notation (HTML output)
   abs  ("\<bar>_\<bar>")
 
+class ord = type +
+  fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  (infix "\<sqsubseteq>" 50)
+    and less :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  (infix "\<sqsubset>" 50)
+begin
+
+notation
+  less_eq  ("op \<^loc><=") and
+  less_eq  ("(_/ \<^loc><= _)" [51, 51] 50) and
+  less  ("op \<^loc><") and
+  less  ("(_/ \<^loc>< _)"  [51, 51] 50)
+  
+notation (xsymbols)
+  less_eq  ("op \<^loc>\<le>") and
+  less_eq  ("(_/ \<^loc>\<le> _)"  [51, 51] 50)
+
+notation (HTML output)
+  less_eq  ("op \<^loc>\<le>") and
+  less_eq  ("(_/ \<^loc>\<le> _)"  [51, 51] 50)
+
+abbreviation (input)
+  greater  (infix "\<^loc>>" 50) where
+  "x \<^loc>> y \<equiv> y \<^loc>< x"
+
+abbreviation (input)
+  greater_eq  (infix "\<^loc>>=" 50) where
+  "x \<^loc>>= y \<equiv> y \<^loc><= x"
+
+notation (input)
+  greater_eq  (infix "\<^loc>\<ge>" 50)
+
+definition
+  Least :: "('a \<Rightarrow> bool) \<Rightarrow> 'a" (binder "\<^loc>LEAST " 10)
+where
+  "Least P == (THE x. P x \<and> (\<forall>y. P y \<longrightarrow> x \<^loc>\<le> y))"
+
+end
+
+notation
+  less_eq  ("op <=") and
+  less_eq  ("(_/ <= _)" [51, 51] 50) and
+  less  ("op <") and
+  less  ("(_/ < _)"  [51, 51] 50)
+  
+notation (xsymbols)
+  less_eq  ("op \<le>") and
+  less_eq  ("(_/ \<le> _)"  [51, 51] 50)
+
+notation (HTML output)
+  less_eq  ("op \<le>") and
+  less_eq  ("(_/ \<le> _)"  [51, 51] 50)
+
+abbreviation (input)
+  greater  (infix ">" 50) where
+  "x > y \<equiv> y < x"
+
+abbreviation (input)
+  greater_eq  (infix ">=" 50) where
+  "x >= y \<equiv> y <= x"
+
+notation (input)
+  greater_eq  (infix "\<ge>" 50)
+
+lemmas Least_def = Least_def [folded ord_class.Least]
+
 syntax
   "_index1"  :: index    ("\<^sub>1")
 translations
--- a/src/HOL/Lattices.thy	Fri Jul 20 00:01:40 2007 +0200
+++ b/src/HOL/Lattices.thy	Fri Jul 20 14:27:56 2007 +0200
@@ -11,12 +11,6 @@
 
 subsection{* Lattices *}
 
-text{*
-  This theory of lattices only defines binary sup and inf
-  operations. The extension to complete lattices is done in theory
-  @{text FixedPoint}.
-*}
-
 class lower_semilattice = order +
   fixes inf :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70)
   assumes inf_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x"
@@ -70,6 +64,9 @@
 
 end
 
+lemma mono_inf: "mono f \<Longrightarrow> f (inf A B) \<le> inf (f A) (f B)"
+  by (auto simp add: mono_def)
+
 
 context upper_semilattice
 begin
@@ -109,6 +106,9 @@
 
 end
 
+lemma mono_sup: "mono f \<Longrightarrow> sup (f A) (f B) \<le> f (sup A B)"
+  by (auto simp add: mono_def)
+
 
 subsubsection{* Equational laws *}
 
@@ -323,6 +323,174 @@
   min_max.le_infI1 min_max.le_infI2
 
 
+subsection {* Complete lattices *}
+
+class complete_lattice = lattice +
+  fixes Inf :: "'a set \<Rightarrow> 'a" ("\<Sqinter>_" [900] 900)
+  assumes Inf_lower: "x \<in> A \<Longrightarrow> \<Sqinter>A \<sqsubseteq> x"
+  assumes Inf_greatest: "(\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> \<Sqinter>A"
+begin
+
+definition
+  Sup :: "'a set \<Rightarrow> 'a" ("\<Squnion>_" [900] 900)
+where
+  "\<Squnion>A = \<Sqinter>{b. \<forall>a \<in> A. a \<^loc>\<le> b}"
+
+lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<^loc>\<le> a}"
+  unfolding Sup_def by (auto intro: Inf_greatest Inf_lower)
+
+lemma Sup_upper: "x \<in> A \<Longrightarrow> x \<^loc>\<le> \<Squnion>A"
+  by (auto simp: Sup_def intro: Inf_greatest)
+
+lemma Sup_least: "(\<And>x. x \<in> A \<Longrightarrow> x \<^loc>\<le> z) \<Longrightarrow> \<Squnion>A \<^loc>\<le> z"
+  by (auto simp: Sup_def intro: Inf_lower)
+
+lemma Inf_Univ: "\<Sqinter>UNIV = \<Squnion>{}"
+  unfolding Sup_def by auto
+
+lemma Sup_Univ: "\<Squnion>UNIV = \<Sqinter>{}"
+  unfolding Inf_Sup by auto
+
+lemma Inf_insert: "\<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
+  apply (rule antisym)
+  apply (rule le_infI)
+  apply (rule Inf_lower)
+  apply simp
+  apply (rule Inf_greatest)
+  apply (rule Inf_lower)
+  apply simp
+  apply (rule Inf_greatest)
+  apply (erule insertE)
+  apply (rule le_infI1)
+  apply simp
+  apply (rule le_infI2)
+  apply (erule Inf_lower)
+  done
+
+lemma Sup_insert: "\<Squnion>insert a A = a \<squnion> \<Squnion>A"
+  apply (rule antisym)
+  apply (rule Sup_least)
+  apply (erule insertE)
+  apply (rule le_supI1)
+  apply simp
+  apply (rule le_supI2)
+  apply (erule Sup_upper)
+  apply (rule le_supI)
+  apply (rule Sup_upper)
+  apply simp
+  apply (rule Sup_least)
+  apply (rule Sup_upper)
+  apply simp
+  done
+
+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)
+
+end
+
+lemmas Sup_def = Sup_def [folded complete_lattice_class.Sup]
+lemmas Sup_upper = Sup_upper [folded complete_lattice_class.Sup]
+lemmas Sup_least = Sup_least [folded complete_lattice_class.Sup]
+
+lemmas Sup_insert [code func] = Sup_insert [folded complete_lattice_class.Sup]
+lemmas Sup_singleton [simp, code func] = Sup_singleton [folded complete_lattice_class.Sup]
+lemmas Sup_insert_simp = Sup_insert_simp [folded complete_lattice_class.Sup]
+lemmas Sup_binary = Sup_binary [folded complete_lattice_class.Sup]
+
+definition
+  top :: "'a::complete_lattice"
+where
+  "top = Inf {}"
+
+definition
+  bot :: "'a::complete_lattice"
+where
+  "bot = Sup {}"
+
+lemma top_greatest [simp]: "x \<le> top"
+  by (unfold top_def, rule Inf_greatest, simp)
+
+lemma bot_least [simp]: "bot \<le> x"
+  by (unfold bot_def, rule Sup_least, simp)
+
+definition
+  SUPR :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b"
+where
+  "SUPR A f == Sup (f ` A)"
+
+definition
+  INFI :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> 'b"
+where
+  "INFI A f == Inf (f ` A)"
+
+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 UNIV (%x. B)"
+  "SUP x. B"     == "SUP x: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 UNIV (%x. B)"
+  "INF x. B"     == "INF x: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
+*}
+
+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: order_antisym SUP_leI le_SUPI)
+
+lemma INF_const[simp]: "A \<noteq> {} \<Longrightarrow> (INF i:A. M) = M"
+  by (auto intro: order_antisym INF_leI le_INFI)
+
+
 subsection {* Bool as lattice *}
 
 instance bool :: distrib_lattice
@@ -330,10 +498,156 @@
   sup_bool_eq: "sup P Q \<equiv> P \<or> Q"
   by intro_classes (auto simp add: inf_bool_eq sup_bool_eq le_bool_def)
 
+instance bool :: complete_lattice
+  Inf_bool_def: "Inf A \<equiv> \<forall>x\<in>A. x"
+  apply intro_classes
+  apply (unfold Inf_bool_def)
+  apply (iprover intro!: le_boolI elim: ballE)
+  apply (iprover intro!: ballI le_boolI elim: ballE le_boolE)
+  done
 
-text {* duplicates *}
+theorem Sup_bool_eq: "Sup A \<longleftrightarrow> (\<exists>x\<in>A. x)"
+  apply (rule order_antisym)
+  apply (rule Sup_least)
+  apply (rule le_boolI)
+  apply (erule bexI, assumption)
+  apply (rule le_boolI)
+  apply (erule bexE)
+  apply (rule le_boolE)
+  apply (rule Sup_upper)
+  apply assumption+
+  done
+
+lemma Inf_empty_bool [simp]:
+  "Inf {}"
+  unfolding Inf_bool_def by auto
+
+lemma not_Sup_empty_bool [simp]:
+  "\<not> Sup {}"
+  unfolding Sup_def Inf_bool_def by auto
+
+lemma top_bool_eq: "top = True"
+  by (iprover intro!: order_antisym le_boolI top_greatest)
+
+lemma bot_bool_eq: "bot = False"
+  by (iprover intro!: order_antisym le_boolI bot_least)
+
+
+subsection {* Set as lattice *}
+
+instance set :: (type) distrib_lattice
+  inf_set_eq: "inf A B \<equiv> A \<inter> B"
+  sup_set_eq: "sup A B \<equiv> A \<union> B"
+  by intro_classes (auto simp add: inf_set_eq sup_set_eq)
+
+lemmas [code func del] = inf_set_eq sup_set_eq
+
+lemmas mono_Int = mono_inf
+  [where ?'a="?'a set", where ?'b="?'b set", unfolded inf_set_eq sup_set_eq]
+
+lemmas mono_Un = mono_sup
+  [where ?'a="?'a set", where ?'b="?'b set", unfolded inf_set_eq sup_set_eq]
+
+instance set :: (type) complete_lattice
+  Inf_set_def: "Inf S \<equiv> \<Inter>S"
+  by intro_classes (auto simp add: Inf_set_def)
+
+lemmas [code func del] = Inf_set_def
+
+theorem Sup_set_eq: "Sup 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)
+
+
+subsection {* Fun as lattice *}
+
+instance "fun" :: (type, lattice) lattice
+  inf_fun_eq: "inf f g \<equiv> (\<lambda>x. inf (f x) (g x))"
+  sup_fun_eq: "sup f g \<equiv> (\<lambda>x. sup (f x) (g x))"
+apply intro_classes
+unfolding inf_fun_eq sup_fun_eq
+apply (auto intro: le_funI)
+apply (rule le_funI)
+apply (auto dest: le_funD)
+apply (rule le_funI)
+apply (auto dest: le_funD)
+done
+
+lemmas [code func del] = inf_fun_eq sup_fun_eq
+
+instance "fun" :: (type, distrib_lattice) distrib_lattice
+  by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1)
+
+instance "fun" :: (type, complete_lattice) complete_lattice
+  Inf_fun_def: "Inf A \<equiv> (\<lambda>x. Inf {y. \<exists>f\<in>A. y = f x})"
+  apply intro_classes
+  apply (unfold Inf_fun_def)
+  apply (rule le_funI)
+  apply (rule Inf_lower)
+  apply (rule CollectI)
+  apply (rule bexI)
+  apply (rule refl)
+  apply assumption
+  apply (rule le_funI)
+  apply (rule Inf_greatest)
+  apply (erule CollectE)
+  apply (erule bexE)
+  apply (iprover elim: le_funE)
+  done
+
+lemmas [code func del] = Inf_fun_def
+
+theorem Sup_fun_eq: "Sup A = (\<lambda>x. Sup {y. \<exists>f\<in>A. y = f x})"
+  apply (rule order_antisym)
+  apply (rule Sup_least)
+  apply (rule le_funI)
+  apply (rule Sup_upper)
+  apply fast
+  apply (rule le_funI)
+  apply (rule Sup_least)
+  apply (erule CollectE)
+  apply (erule bexE)
+  apply (drule le_funD [OF Sup_upper])
+  apply simp
+  done
+
+lemma Inf_empty_fun:
+  "Inf {} = (\<lambda>_. Inf {})"
+  by rule (auto simp add: Inf_fun_def)
+
+lemma Sup_empty_fun:
+  "Sup {} = (\<lambda>_. Sup {})"
+proof -
+  have aux: "\<And>x. {y. \<exists>f. y = f x} = UNIV" by auto
+  show ?thesis
+  by (auto simp add: Sup_def Inf_fun_def Inf_binary inf_bool_eq aux)
+qed
+
+lemma top_fun_eq: "top = (\<lambda>x. top)"
+  by (iprover intro!: order_antisym le_funI top_greatest)
+
+lemma bot_fun_eq: "bot = (\<lambda>x. bot)"
+  by (iprover intro!: order_antisym le_funI bot_least)
+
+
+text {* redundant bindings *}
 
 lemmas inf_aci = inf_ACI
 lemmas sup_aci = sup_ACI
 
+ML {*
+val sup_fun_eq = @{thm sup_fun_eq}
+val sup_bool_eq = @{thm sup_bool_eq}
+*}
+
 end
--- a/src/HOL/Set.thy	Fri Jul 20 00:01:40 2007 +0200
+++ b/src/HOL/Set.thy	Fri Jul 20 14:27:56 2007 +0200
@@ -6,7 +6,7 @@
 header {* Set theory for higher-order logic *}
 
 theory Set
-imports Lattices
+imports HOL
 begin
 
 text {* A set in HOL is simply a predicate. *}
@@ -1040,13 +1040,6 @@
   and [symmetric, defn] = atomize_ball
 
 
-subsection {* Order on sets *}
-
-instance set :: (type) order
-  by (intro_classes,
-      (assumption | rule subset_refl subset_trans subset_antisym psubset_eq)+)
-
-
 subsection {* Further set-theory lemmas *}
 
 subsubsection {* Derived rules involving subsets. *}
@@ -1054,12 +1047,10 @@
 text {* @{text insert}. *}
 
 lemma subset_insertI: "B \<subseteq> insert a B"
-  apply (rule subsetI)
-  apply (erule insertI2)
-  done
+  by (rule subsetI) (erule insertI2)
 
 lemma subset_insertI2: "A \<subseteq> B \<Longrightarrow> A \<subseteq> insert b B"
-by blast
+  by blast
 
 lemma subset_insert: "x \<notin> A ==> (A \<subseteq> insert x B) = (A \<subseteq> B)"
   by blast
@@ -1135,14 +1126,6 @@
 by blast
 
 
-text {* \medskip Monotonicity. *}
-
-lemma mono_Un: "mono f ==> f A \<union> f B \<subseteq> f (A \<union> B)"
-  by (auto simp add: mono_def)
-
-lemma mono_Int: "mono f ==> f (A \<inter> B) \<subseteq> f A \<inter> f B"
-  by (auto simp add: mono_def)
-
 subsubsection {* Equalities involving union, intersection, inclusion, etc. *}
 
 text {* @{text "{}"}. *}
@@ -2014,16 +1997,6 @@
 lemma eq_to_mono2: "a = b ==> c = d ==> ~ b --> ~ d ==> ~ a --> ~ c"
   by iprover
 
-lemma Least_mono:
-  "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y
-    ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)"
-    -- {* Courtesy of Stephan Merz *}
-  apply clarify
-  apply (erule_tac P = "%x. x : S" in LeastI2_order, fast)
-  apply (rule LeastI2_order)
-  apply (auto elim: monoD intro!: order_antisym)
-  done
-
 
 subsection {* Inverse image of a function *}
 
@@ -2120,19 +2093,6 @@
 lemma set_mp: "A \<subseteq> B ==> x:A ==> x:B"
   by (rule subsetD)
 
-lemmas basic_trans_rules [trans] =
-  order_trans_rules set_rev_mp set_mp
-
-
-subsection {* Sets as lattice *}
-
-instance set :: (type) distrib_lattice
-  inf_set_eq: "inf A B \<equiv> A \<inter> B"
-  sup_set_eq: "sup A B \<equiv> A \<union> B"
-  by intro_classes (auto simp add: inf_set_eq sup_set_eq)
-
-lemmas [code func del] = inf_set_eq sup_set_eq
-
 
 subsection {* Basic ML bindings *}