src/HOL/Conditionally_Complete_Lattices.thy
changeset 55711 71c701dc5bf9
parent 55710 adfc759263ab
child 55713 89991ef58448
     1.1 --- a/src/HOL/Conditionally_Complete_Lattices.thy	Tue Nov 05 09:44:58 2013 +0100
     1.2 +++ b/src/HOL/Conditionally_Complete_Lattices.thy	Tue Nov 05 09:44:59 2013 +0100
     1.3 @@ -10,10 +10,10 @@
     1.4  imports Main Lubs
     1.5  begin
     1.6  
     1.7 -lemma Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
     1.8 +lemma (in linorder) Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
     1.9    by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
    1.10  
    1.11 -lemma Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
    1.12 +lemma (in linorder) Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
    1.13    by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
    1.14  
    1.15  context preorder
    1.16 @@ -125,6 +125,12 @@
    1.17    thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
    1.18  qed
    1.19  
    1.20 +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)"
    1.21 +  by (auto simp: bdd_above_def intro: le_supI1 le_supI2)
    1.22 +
    1.23 +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)"
    1.24 +  by (auto simp: bdd_below_def intro: le_infI1 le_infI2)
    1.25 +
    1.26  end
    1.27  
    1.28  
    1.29 @@ -142,6 +148,24 @@
    1.30      and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
    1.31  begin
    1.32  
    1.33 +lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
    1.34 +  by (metis cSup_upper order_trans)
    1.35 +
    1.36 +lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
    1.37 +  by (metis cInf_lower order_trans)
    1.38 +
    1.39 +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"
    1.40 +  by (metis cSup_least cSup_upper2)
    1.41 +
    1.42 +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"
    1.43 +  by (metis cInf_greatest cInf_lower2)
    1.44 +
    1.45 +lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
    1.46 +  by (metis cSup_least cSup_upper subsetD)
    1.47 +
    1.48 +lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
    1.49 +  by (metis cInf_greatest cInf_lower subsetD)
    1.50 +
    1.51  lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
    1.52    by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
    1.53  
    1.54 @@ -154,18 +178,6 @@
    1.55  lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
    1.56    by (metis order_trans cInf_lower cInf_greatest)
    1.57  
    1.58 -lemma cSup_singleton [simp]: "Sup {x} = x"
    1.59 -  by (intro cSup_eq_maximum) auto
    1.60 -
    1.61 -lemma cInf_singleton [simp]: "Inf {x} = x"
    1.62 -  by (intro cInf_eq_minimum) auto
    1.63 -
    1.64 -lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
    1.65 -  by (metis cSup_upper order_trans)
    1.66 - 
    1.67 -lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
    1.68 -  by (metis cInf_lower order_trans)
    1.69 -
    1.70  lemma cSup_eq_non_empty:
    1.71    assumes 1: "X \<noteq> {}"
    1.72    assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
    1.73 @@ -192,10 +204,16 @@
    1.74  lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
    1.75    by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
    1.76  
    1.77 +lemma cSup_singleton [simp]: "Sup {x} = x"
    1.78 +  by (intro cSup_eq_maximum) auto
    1.79 +
    1.80 +lemma cInf_singleton [simp]: "Inf {x} = x"
    1.81 +  by (intro cInf_eq_minimum) auto
    1.82 +
    1.83  lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
    1.84    using cSup_insert[of X] by simp
    1.85  
    1.86 -lemma cInf_insert_if: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
    1.87 +lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
    1.88    using cInf_insert[of X] by simp
    1.89  
    1.90  lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
    1.91 @@ -234,6 +252,74 @@
    1.92  lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
    1.93    by (auto intro!: cInf_eq_minimum)
    1.94  
    1.95 +lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> INFI A f \<le> f x"
    1.96 +  unfolding INF_def by (rule cInf_lower) auto
    1.97 +
    1.98 +lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> INFI A f"
    1.99 +  unfolding INF_def by (rule cInf_greatest) auto
   1.100 +
   1.101 +lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> SUPR A f"
   1.102 +  unfolding SUP_def by (rule cSup_upper) auto
   1.103 +
   1.104 +lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> SUPR A f \<le> M"
   1.105 +  unfolding SUP_def by (rule cSup_least) auto
   1.106 +
   1.107 +lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> INFI A f \<le> u"
   1.108 +  by (auto intro: cINF_lower assms order_trans)
   1.109 +
   1.110 +lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> SUPR A f"
   1.111 +  by (auto intro: cSUP_upper assms order_trans)
   1.112 +
   1.113 +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)"
   1.114 +  by (metis cINF_greatest cINF_lower assms order_trans)
   1.115 +
   1.116 +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)"
   1.117 +  by (metis cSUP_least cSUP_upper assms order_trans)
   1.118 +
   1.119 +lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> INFI (insert a A) f = inf (f a) (INFI A f)"
   1.120 +  by (metis INF_def cInf_insert assms empty_is_image image_insert)
   1.121 +
   1.122 +lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR (insert a A) f = sup (f a) (SUPR A f)"
   1.123 +  by (metis SUP_def cSup_insert assms empty_is_image image_insert)
   1.124 +
   1.125 +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"
   1.126 +  unfolding INF_def by (auto intro: cInf_mono)
   1.127 +
   1.128 +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"
   1.129 +  unfolding SUP_def by (auto intro: cSup_mono)
   1.130 +
   1.131 +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"
   1.132 +  by (rule cINF_mono) auto
   1.133 +
   1.134 +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"
   1.135 +  by (rule cSUP_mono) auto
   1.136 +
   1.137 +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)"
   1.138 +  by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1)
   1.139 +
   1.140 +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) "
   1.141 +  by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
   1.142 +
   1.143 +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)"
   1.144 +  by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
   1.145 +
   1.146 +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)"
   1.147 +  unfolding INF_def using assms by (auto simp add: image_Un intro: cInf_union_distrib)
   1.148 +
   1.149 +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)"
   1.150 +  by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
   1.151 +
   1.152 +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)"
   1.153 +  unfolding SUP_def by (auto simp add: image_Un intro: cSup_union_distrib)
   1.154 +
   1.155 +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))"
   1.156 +  by (intro antisym le_infI cINF_greatest cINF_lower2)
   1.157 +     (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
   1.158 +
   1.159 +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))"
   1.160 +  by (intro antisym le_supI cSUP_least cSUP_upper2)
   1.161 +     (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
   1.162 +
   1.163  end
   1.164  
   1.165  instance complete_lattice \<subseteq> conditionally_complete_lattice
   1.166 @@ -323,14 +409,11 @@
   1.167  
   1.168  end
   1.169  
   1.170 -class linear_continuum = conditionally_complete_linorder + dense_linorder +
   1.171 -  assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
   1.172 -begin
   1.173 +lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
   1.174 +  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
   1.175  
   1.176 -lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
   1.177 -  by (metis UNIV_not_singleton neq_iff)
   1.178 -
   1.179 -end
   1.180 +lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
   1.181 +  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
   1.182  
   1.183  lemma cSup_bounds:
   1.184    fixes S :: "'a :: conditionally_complete_lattice set"
   1.185 @@ -347,19 +430,12 @@
   1.186    with b show ?thesis by blast
   1.187  qed
   1.188  
   1.189 -
   1.190  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"
   1.191    by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
   1.192  
   1.193  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"
   1.194    by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
   1.195  
   1.196 -lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
   1.197 -  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
   1.198 -
   1.199 -lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
   1.200 -  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
   1.201 -
   1.202  lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
   1.203    by (auto intro!: cSup_eq_non_empty intro: dense_le)
   1.204  
   1.205 @@ -378,4 +454,13 @@
   1.206  lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
   1.207    by (auto intro!: cInf_eq intro: dense_ge_bounded)
   1.208  
   1.209 +class linear_continuum = conditionally_complete_linorder + dense_linorder +
   1.210 +  assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
   1.211 +begin
   1.212 +
   1.213 +lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
   1.214 +  by (metis UNIV_not_singleton neq_iff)
   1.215 +
   1.216  end
   1.217 +
   1.218 +end