--- 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