--- a/src/HOL/Complete_Lattices.thy Tue Nov 05 09:44:58 2013 +0100
+++ b/src/HOL/Complete_Lattices.thy Tue Nov 05 09:44:59 2013 +0100
@@ -20,6 +20,12 @@
definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
INF_def: "INFI A f = \<Sqinter>(f ` A)"
+lemma INF_image [simp]: "INFI (f`A) g = INFI A (\<lambda>x. g (f x))"
+ by (simp add: INF_def image_image)
+
+lemma INF_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> INFI A C = INFI B D"
+ by (simp add: INF_def image_def)
+
end
class Sup =
@@ -29,6 +35,12 @@
definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
SUP_def: "SUPR A f = \<Squnion>(f ` A)"
+lemma SUP_image [simp]: "SUPR (f`A) g = SUPR A (%x. g (f x))"
+ by (simp add: SUP_def image_image)
+
+lemma SUP_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> SUPR A C = SUPR B D"
+ by (simp add: SUP_def image_def)
+
end
text {*
@@ -183,12 +195,6 @@
"\<Squnion>UNIV = \<top>"
by (auto intro!: antisym Sup_upper)
-lemma INF_image [simp]: "(\<Sqinter>x\<in>f`A. g x) = (\<Sqinter>x\<in>A. g (f x))"
- by (simp add: INF_def image_image)
-
-lemma SUP_image [simp]: "(\<Squnion>x\<in>f`A. g x) = (\<Squnion>x\<in>A. g (f x))"
- by (simp add: SUP_def image_image)
-
lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<sqsubseteq> a}"
by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
@@ -201,14 +207,6 @@
lemma Sup_subset_mono: "A \<subseteq> B \<Longrightarrow> \<Squnion>A \<sqsubseteq> \<Squnion>B"
by (auto intro: Sup_least Sup_upper)
-lemma INF_cong:
- "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)"
- by (simp add: INF_def image_def)
-
-lemma SUP_cong:
- "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)"
- by (simp add: SUP_def image_def)
-
lemma Inf_mono:
assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<sqsubseteq> b"
shows "\<Sqinter>A \<sqsubseteq> \<Sqinter>B"
--- a/src/HOL/Conditionally_Complete_Lattices.thy Tue Nov 05 09:44:58 2013 +0100
+++ b/src/HOL/Conditionally_Complete_Lattices.thy Tue Nov 05 09:44:59 2013 +0100
@@ -10,10 +10,10 @@
imports Main Lubs
begin
-lemma Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
+lemma (in linorder) Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
-lemma Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
+lemma (in linorder) Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
context preorder
@@ -125,6 +125,12 @@
thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
qed
+lemma bdd_above_sup[simp]: "bdd_above ((\<lambda>x. sup (f x) (g x)) ` A) \<longleftrightarrow> bdd_above (f`A) \<and> bdd_above (g`A)"
+ by (auto simp: bdd_above_def intro: le_supI1 le_supI2)
+
+lemma bdd_below_inf[simp]: "bdd_below ((\<lambda>x. inf (f x) (g x)) ` A) \<longleftrightarrow> bdd_below (f`A) \<and> bdd_below (g`A)"
+ by (auto simp: bdd_below_def intro: le_infI1 le_infI2)
+
end
@@ -142,6 +148,24 @@
and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
begin
+lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
+ by (metis cSup_upper order_trans)
+
+lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
+ by (metis cInf_lower order_trans)
+
+lemma cSup_mono: "B \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. b \<le> a) \<Longrightarrow> Sup B \<le> Sup A"
+ by (metis cSup_least cSup_upper2)
+
+lemma cInf_mono: "B \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b) \<Longrightarrow> Inf A \<le> Inf B"
+ by (metis cInf_greatest cInf_lower2)
+
+lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
+ by (metis cSup_least cSup_upper subsetD)
+
+lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
+ by (metis cInf_greatest cInf_lower subsetD)
+
lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
@@ -154,18 +178,6 @@
lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
by (metis order_trans cInf_lower cInf_greatest)
-lemma cSup_singleton [simp]: "Sup {x} = x"
- by (intro cSup_eq_maximum) auto
-
-lemma cInf_singleton [simp]: "Inf {x} = x"
- by (intro cInf_eq_minimum) auto
-
-lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
- by (metis cSup_upper order_trans)
-
-lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
- by (metis cInf_lower order_trans)
-
lemma cSup_eq_non_empty:
assumes 1: "X \<noteq> {}"
assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
@@ -192,10 +204,16 @@
lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
+lemma cSup_singleton [simp]: "Sup {x} = x"
+ by (intro cSup_eq_maximum) auto
+
+lemma cInf_singleton [simp]: "Inf {x} = x"
+ by (intro cInf_eq_minimum) auto
+
lemma cSup_insert_If: "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
using cSup_insert[of X] by simp
-lemma cInf_insert_if: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
+lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
using cInf_insert[of X] by simp
lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
@@ -234,6 +252,74 @@
lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
by (auto intro!: cInf_eq_minimum)
+lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> INFI A f \<le> f x"
+ unfolding INF_def by (rule cInf_lower) auto
+
+lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> INFI A f"
+ unfolding INF_def by (rule cInf_greatest) auto
+
+lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> SUPR A f"
+ unfolding SUP_def by (rule cSup_upper) auto
+
+lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> SUPR A f \<le> M"
+ unfolding SUP_def by (rule cSup_least) auto
+
+lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> INFI A f \<le> u"
+ by (auto intro: cINF_lower assms order_trans)
+
+lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> SUPR A f"
+ by (auto intro: cSUP_upper assms order_trans)
+
+lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> INFI A f \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
+ by (metis cINF_greatest cINF_lower assms order_trans)
+
+lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR A f \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
+ by (metis cSUP_least cSUP_upper assms order_trans)
+
+lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> INFI (insert a A) f = inf (f a) (INFI A f)"
+ by (metis INF_def cInf_insert assms empty_is_image image_insert)
+
+lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR (insert a A) f = sup (f a) (SUPR A f)"
+ by (metis SUP_def cSup_insert assms empty_is_image image_insert)
+
+lemma cINF_mono: "B \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> INFI A f \<le> INFI B g"
+ unfolding INF_def by (auto intro: cInf_mono)
+
+lemma cSUP_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> (\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> SUPR A f \<le> SUPR B g"
+ unfolding SUP_def by (auto intro: cSup_mono)
+
+lemma cINF_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> g x \<le> f x) \<Longrightarrow> INFI B g \<le> INFI A f"
+ by (rule cINF_mono) auto
+
+lemma cSUP_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> SUPR A f \<le> SUPR B g"
+ by (rule cSUP_mono) auto
+
+lemma less_eq_cInf_inter: "bdd_below A \<Longrightarrow> bdd_below B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> inf (Inf A) (Inf B) \<le> Inf (A \<inter> B)"
+ by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1)
+
+lemma cSup_inter_less_eq: "bdd_above A \<Longrightarrow> bdd_above B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> Sup (A \<inter> B) \<le> sup (Sup A) (Sup B) "
+ by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
+
+lemma cInf_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> Inf (A \<union> B) = inf (Inf A) (Inf B)"
+ by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
+
+lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f`B) \<Longrightarrow> INFI (A \<union> B) f = inf (INFI A f) (INFI B f)"
+ unfolding INF_def using assms by (auto simp add: image_Un intro: cInf_union_distrib)
+
+lemma cSup_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> Sup (A \<union> B) = sup (Sup A) (Sup B)"
+ by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
+
+lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f`B) \<Longrightarrow> SUPR (A \<union> B) f = sup (SUPR A f) (SUPR B f)"
+ unfolding SUP_def by (auto simp add: image_Un intro: cSup_union_distrib)
+
+lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> inf (INFI A f) (INFI A g) = (INF a:A. inf (f a) (g a))"
+ by (intro antisym le_infI cINF_greatest cINF_lower2)
+ (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
+
+lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> sup (SUPR A f) (SUPR A g) = (SUP a:A. sup (f a) (g a))"
+ by (intro antisym le_supI cSUP_least cSUP_upper2)
+ (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
+
end
instance complete_lattice \<subseteq> conditionally_complete_lattice
@@ -323,14 +409,11 @@
end
-class linear_continuum = conditionally_complete_linorder + dense_linorder +
- assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
-begin
+lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
+ using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
-lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
- by (metis UNIV_not_singleton neq_iff)
-
-end
+lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
+ using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
lemma cSup_bounds:
fixes S :: "'a :: conditionally_complete_lattice set"
@@ -347,19 +430,12 @@
with b show ?thesis by blast
qed
-
lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
-lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
- using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
-
-lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
- using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
-
lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
by (auto intro!: cSup_eq_non_empty intro: dense_le)
@@ -378,4 +454,13 @@
lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
by (auto intro!: cInf_eq intro: dense_ge_bounded)
+class linear_continuum = conditionally_complete_linorder + dense_linorder +
+ assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
+begin
+
+lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
+ by (metis UNIV_not_singleton neq_iff)
+
end
+
+end
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Nov 05 09:44:58 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Nov 05 09:44:59 2013 +0100
@@ -2686,7 +2686,7 @@
lemma Inf_insert:
fixes S :: "real set"
shows "bounded S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
- by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_if)
+ by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If)
lemma Inf_insert_finite:
fixes S :: "real set"