| author | haftmann | 
| Mon, 09 Jan 2017 18:53:20 +0100 | |
| changeset 64849 | 766db3539859 | 
| parent 62390 | 842917225d56 | 
| child 69260 | 0a9688695a1b | 
| permissions | -rw-r--r-- | 
| 62373 | 1 | (* Title: HOL/Library/Countable_Complete_Lattices.thy | 
| 2 | Author: Johannes Hölzl | |
| 3 | *) | |
| 4 | ||
| 5 | section \<open>Countable Complete Lattices\<close> | |
| 6 | ||
| 7 | theory Countable_Complete_Lattices | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 8 | imports Main Countable_Set | 
| 62373 | 9 | begin | 
| 10 | ||
| 11 | lemma UNIV_nat_eq: "UNIV = insert 0 (range Suc)" | |
| 12 | by (metis UNIV_eq_I nat.nchotomy insertCI rangeI) | |
| 13 | ||
| 14 | class countable_complete_lattice = lattice + Inf + Sup + bot + top + | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 15 | assumes ccInf_lower: "countable A \<Longrightarrow> x \<in> A \<Longrightarrow> Inf A \<le> x" | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 16 | assumes ccInf_greatest: "countable A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf A" | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 17 | assumes ccSup_upper: "countable A \<Longrightarrow> x \<in> A \<Longrightarrow> x \<le> Sup A" | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 18 | assumes ccSup_least: "countable A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup A \<le> z" | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 19 |   assumes ccInf_empty [simp]: "Inf {} = top"
 | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 20 |   assumes ccSup_empty [simp]: "Sup {} = bot"
 | 
| 62373 | 21 | begin | 
| 22 | ||
| 23 | subclass bounded_lattice | |
| 24 | proof | |
| 25 | fix a | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 26 | show "bot \<le> a" by (auto intro: ccSup_least simp only: ccSup_empty [symmetric]) | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 27 | show "a \<le> top" by (auto intro: ccInf_greatest simp only: ccInf_empty [symmetric]) | 
| 62373 | 28 | qed | 
| 29 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 30 | lemma ccINF_lower: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> (INF i :A. f i) \<le> f i" | 
| 62373 | 31 | using ccInf_lower [of "f ` A"] by simp | 
| 32 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 33 | lemma ccINF_greatest: "countable A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> u \<le> f i) \<Longrightarrow> u \<le> (INF i :A. f i)" | 
| 62373 | 34 | using ccInf_greatest [of "f ` A"] by auto | 
| 35 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 36 | lemma ccSUP_upper: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> f i \<le> (SUP i :A. f i)" | 
| 62373 | 37 | using ccSup_upper [of "f ` A"] by simp | 
| 38 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 39 | lemma ccSUP_least: "countable A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> f i \<le> u) \<Longrightarrow> (SUP i :A. f i) \<le> u" | 
| 62373 | 40 | using ccSup_least [of "f ` A"] by auto | 
| 41 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 42 | lemma ccInf_lower2: "countable A \<Longrightarrow> u \<in> A \<Longrightarrow> u \<le> v \<Longrightarrow> Inf A \<le> v" | 
| 62373 | 43 | using ccInf_lower [of A u] by auto | 
| 44 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 45 | lemma ccINF_lower2: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> f i \<le> u \<Longrightarrow> (INF i :A. f i) \<le> u" | 
| 62373 | 46 | using ccINF_lower [of A i f] by auto | 
| 47 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 48 | lemma ccSup_upper2: "countable A \<Longrightarrow> u \<in> A \<Longrightarrow> v \<le> u \<Longrightarrow> v \<le> Sup A" | 
| 62373 | 49 | using ccSup_upper [of A u] by auto | 
| 50 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 51 | lemma ccSUP_upper2: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> u \<le> f i \<Longrightarrow> u \<le> (SUP i :A. f i)" | 
| 62373 | 52 | using ccSUP_upper [of A i f] by auto | 
| 53 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 54 | lemma le_ccInf_iff: "countable A \<Longrightarrow> b \<le> Inf A \<longleftrightarrow> (\<forall>a\<in>A. b \<le> a)" | 
| 62373 | 55 | by (auto intro: ccInf_greatest dest: ccInf_lower) | 
| 56 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 57 | lemma le_ccINF_iff: "countable A \<Longrightarrow> u \<le> (INF i :A. f i) \<longleftrightarrow> (\<forall>i\<in>A. u \<le> f i)" | 
| 62373 | 58 | using le_ccInf_iff [of "f ` A"] by simp | 
| 59 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 60 | lemma ccSup_le_iff: "countable A \<Longrightarrow> Sup A \<le> b \<longleftrightarrow> (\<forall>a\<in>A. a \<le> b)" | 
| 62373 | 61 | by (auto intro: ccSup_least dest: ccSup_upper) | 
| 62 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 63 | lemma ccSUP_le_iff: "countable A \<Longrightarrow> (SUP i :A. f i) \<le> u \<longleftrightarrow> (\<forall>i\<in>A. f i \<le> u)" | 
| 62373 | 64 | using ccSup_le_iff [of "f ` A"] by simp | 
| 65 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 66 | lemma ccInf_insert [simp]: "countable A \<Longrightarrow> Inf (insert a A) = inf a (Inf A)" | 
| 62373 | 67 | by (force intro: le_infI le_infI1 le_infI2 antisym ccInf_greatest ccInf_lower) | 
| 68 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 69 | lemma ccINF_insert [simp]: "countable A \<Longrightarrow> (INF x:insert a A. f x) = inf (f a) (INFIMUM A f)" | 
| 62373 | 70 | unfolding image_insert by simp | 
| 71 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 72 | lemma ccSup_insert [simp]: "countable A \<Longrightarrow> Sup (insert a A) = sup a (Sup A)" | 
| 62373 | 73 | by (force intro: le_supI le_supI1 le_supI2 antisym ccSup_least ccSup_upper) | 
| 74 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 75 | lemma ccSUP_insert [simp]: "countable A \<Longrightarrow> (SUP x:insert a A. f x) = sup (f a) (SUPREMUM A f)" | 
| 62373 | 76 | unfolding image_insert by simp | 
| 77 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 78 | lemma ccINF_empty [simp]: "(INF x:{}. f x) = top"
 | 
| 62373 | 79 | unfolding image_empty by simp | 
| 80 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 81 | lemma ccSUP_empty [simp]: "(SUP x:{}. f x) = bot"
 | 
| 62373 | 82 | unfolding image_empty by simp | 
| 83 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 84 | lemma ccInf_superset_mono: "countable A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> Inf A \<le> Inf B" | 
| 62373 | 85 | by (auto intro: ccInf_greatest ccInf_lower countable_subset) | 
| 86 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 87 | lemma ccSup_subset_mono: "countable B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B" | 
| 62373 | 88 | by (auto intro: ccSup_least ccSup_upper countable_subset) | 
| 89 | ||
| 90 | lemma ccInf_mono: | |
| 91 | assumes [intro]: "countable B" "countable A" | |
| 92 | assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b" | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 93 | shows "Inf A \<le> Inf B" | 
| 62373 | 94 | proof (rule ccInf_greatest) | 
| 95 | fix b assume "b \<in> B" | |
| 96 | with assms obtain a where "a \<in> A" and "a \<le> b" by blast | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 97 | from \<open>a \<in> A\<close> have "Inf A \<le> a" by (rule ccInf_lower[rotated]) auto | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 98 | with \<open>a \<le> b\<close> show "Inf A \<le> b" by auto | 
| 62373 | 99 | qed auto | 
| 100 | ||
| 101 | lemma ccINF_mono: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 102 | "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> (INF n:A. f n) \<le> (INF n:B. g n)" | 
| 62373 | 103 | using ccInf_mono [of "g ` B" "f ` A"] by auto | 
| 104 | ||
| 105 | lemma ccSup_mono: | |
| 106 | assumes [intro]: "countable B" "countable A" | |
| 107 | assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<le> b" | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 108 | shows "Sup A \<le> Sup B" | 
| 62373 | 109 | proof (rule ccSup_least) | 
| 110 | fix a assume "a \<in> A" | |
| 111 | with assms obtain b where "b \<in> B" and "a \<le> b" by blast | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 112 | from \<open>b \<in> B\<close> have "b \<le> Sup B" by (rule ccSup_upper[rotated]) auto | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 113 | with \<open>a \<le> b\<close> show "a \<le> Sup B" by auto | 
| 62373 | 114 | qed auto | 
| 115 | ||
| 116 | lemma ccSUP_mono: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 117 | "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> (SUP n:A. f n) \<le> (SUP n:B. g n)" | 
| 62373 | 118 | using ccSup_mono [of "g ` B" "f ` A"] by auto | 
| 119 | ||
| 120 | lemma ccINF_superset_mono: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 121 | "countable A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (INF x:A. f x) \<le> (INF x:B. g x)" | 
| 62373 | 122 | by (blast intro: ccINF_mono countable_subset dest: subsetD) | 
| 123 | ||
| 124 | lemma ccSUP_subset_mono: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 125 | "countable B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (SUP x:A. f x) \<le> (SUP x:B. g x)" | 
| 62373 | 126 | by (blast intro: ccSUP_mono countable_subset dest: subsetD) | 
| 127 | ||
| 128 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 129 | lemma less_eq_ccInf_inter: "countable A \<Longrightarrow> countable B \<Longrightarrow> sup (Inf A) (Inf B) \<le> Inf (A \<inter> B)" | 
| 62373 | 130 | by (auto intro: ccInf_greatest ccInf_lower) | 
| 131 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 132 | lemma ccSup_inter_less_eq: "countable A \<Longrightarrow> countable B \<Longrightarrow> Sup (A \<inter> B) \<le> inf (Sup A) (Sup B)" | 
| 62373 | 133 | by (auto intro: ccSup_least ccSup_upper) | 
| 134 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 135 | lemma ccInf_union_distrib: "countable A \<Longrightarrow> countable B \<Longrightarrow> Inf (A \<union> B) = inf (Inf A) (Inf B)" | 
| 62373 | 136 | by (rule antisym) (auto intro: ccInf_greatest ccInf_lower le_infI1 le_infI2) | 
| 137 | ||
| 138 | lemma ccINF_union: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 139 | "countable A \<Longrightarrow> countable B \<Longrightarrow> (INF i:A \<union> B. M i) = inf (INF i:A. M i) (INF i:B. M i)" | 
| 62373 | 140 | by (auto intro!: antisym ccINF_mono intro: le_infI1 le_infI2 ccINF_greatest ccINF_lower) | 
| 141 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 142 | lemma ccSup_union_distrib: "countable A \<Longrightarrow> countable B \<Longrightarrow> Sup (A \<union> B) = sup (Sup A) (Sup B)" | 
| 62373 | 143 | by (rule antisym) (auto intro: ccSup_least ccSup_upper le_supI1 le_supI2) | 
| 144 | ||
| 145 | lemma ccSUP_union: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 146 | "countable A \<Longrightarrow> countable B \<Longrightarrow> (SUP i:A \<union> B. M i) = sup (SUP i:A. M i) (SUP i:B. M i)" | 
| 62373 | 147 | by (auto intro!: antisym ccSUP_mono intro: le_supI1 le_supI2 ccSUP_least ccSUP_upper) | 
| 148 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 149 | lemma ccINF_inf_distrib: "countable A \<Longrightarrow> inf (INF a:A. f a) (INF a:A. g a) = (INF a:A. inf (f a) (g a))" | 
| 62373 | 150 | by (rule antisym) (rule ccINF_greatest, auto intro: le_infI1 le_infI2 ccINF_lower ccINF_mono) | 
| 151 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 152 | lemma ccSUP_sup_distrib: "countable A \<Longrightarrow> sup (SUP a:A. f a) (SUP a:A. g a) = (SUP a:A. sup (f a) (g a))" | 
| 62373 | 153 | by (rule antisym[rotated]) (rule ccSUP_least, auto intro: le_supI1 le_supI2 ccSUP_upper ccSUP_mono) | 
| 154 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 155 | lemma ccINF_const [simp]: "A \<noteq> {} \<Longrightarrow> (INF i :A. f) = f"
 | 
| 62373 | 156 | unfolding image_constant_conv by auto | 
| 157 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 158 | lemma ccSUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (SUP i :A. f) = f"
 | 
| 62373 | 159 | unfolding image_constant_conv by auto | 
| 160 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 161 | lemma ccINF_top [simp]: "(INF x:A. top) = top" | 
| 62373 | 162 |   by (cases "A = {}") simp_all
 | 
| 163 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 164 | lemma ccSUP_bot [simp]: "(SUP x:A. bot) = bot" | 
| 62373 | 165 |   by (cases "A = {}") simp_all
 | 
| 166 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 167 | lemma ccINF_commute: "countable A \<Longrightarrow> countable B \<Longrightarrow> (INF i:A. INF j:B. f i j) = (INF j:B. INF i:A. f i j)" | 
| 62373 | 168 | by (iprover intro: ccINF_lower ccINF_greatest order_trans antisym) | 
| 169 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 170 | lemma ccSUP_commute: "countable A \<Longrightarrow> countable B \<Longrightarrow> (SUP i:A. SUP j:B. f i j) = (SUP j:B. SUP i:A. f i j)" | 
| 62373 | 171 | by (iprover intro: ccSUP_upper ccSUP_least order_trans antisym) | 
| 172 | ||
| 173 | end | |
| 174 | ||
| 175 | context | |
| 176 |   fixes a :: "'a::{countable_complete_lattice, linorder}"
 | |
| 177 | begin | |
| 178 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 179 | lemma less_ccSup_iff: "countable S \<Longrightarrow> a < Sup S \<longleftrightarrow> (\<exists>x\<in>S. a < x)" | 
| 62373 | 180 | unfolding not_le [symmetric] by (subst ccSup_le_iff) auto | 
| 181 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 182 | lemma less_ccSUP_iff: "countable A \<Longrightarrow> a < (SUP i:A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)" | 
| 62373 | 183 | using less_ccSup_iff [of "f ` A"] by simp | 
| 184 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 185 | lemma ccInf_less_iff: "countable S \<Longrightarrow> Inf S < a \<longleftrightarrow> (\<exists>x\<in>S. x < a)" | 
| 62373 | 186 | unfolding not_le [symmetric] by (subst le_ccInf_iff) auto | 
| 187 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 188 | lemma ccINF_less_iff: "countable A \<Longrightarrow> (INF i:A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)" | 
| 62373 | 189 | using ccInf_less_iff [of "f ` A"] by simp | 
| 190 | ||
| 191 | end | |
| 192 | ||
| 193 | class countable_complete_distrib_lattice = countable_complete_lattice + | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 194 | assumes sup_ccInf: "countable B \<Longrightarrow> sup a (Inf B) = (INF b:B. sup a b)" | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 195 | assumes inf_ccSup: "countable B \<Longrightarrow> inf a (Sup B) = (SUP b:B. inf a b)" | 
| 62373 | 196 | begin | 
| 197 | ||
| 198 | lemma sup_ccINF: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 199 | "countable B \<Longrightarrow> sup a (INF b:B. f b) = (INF b:B. sup a (f b))" | 
| 62373 | 200 | by (simp only: sup_ccInf image_image countable_image) | 
| 201 | ||
| 202 | lemma inf_ccSUP: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 203 | "countable B \<Longrightarrow> inf a (SUP b:B. f b) = (SUP b:B. inf a (f b))" | 
| 62373 | 204 | by (simp only: inf_ccSup image_image countable_image) | 
| 205 | ||
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 206 | subclass distrib_lattice | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 207 | proof | 
| 62373 | 208 | fix a b c | 
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 209 |   from sup_ccInf[of "{b, c}" a] have "sup a (Inf {b, c}) = (INF d:{b, c}. sup a d)"
 | 
| 62373 | 210 | by simp | 
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 211 | then show "sup a (inf b c) = inf (sup a b) (sup a c)" | 
| 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 212 | by simp | 
| 62373 | 213 | qed | 
| 214 | ||
| 215 | lemma ccInf_sup: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 216 | "countable B \<Longrightarrow> sup (Inf B) a = (INF b:B. sup b a)" | 
| 62373 | 217 | by (simp add: sup_ccInf sup_commute) | 
| 218 | ||
| 219 | lemma ccSup_inf: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 220 | "countable B \<Longrightarrow> inf (Sup B) a = (SUP b:B. inf b a)" | 
| 62373 | 221 | by (simp add: inf_ccSup inf_commute) | 
| 222 | ||
| 223 | lemma ccINF_sup: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 224 | "countable B \<Longrightarrow> sup (INF b:B. f b) a = (INF b:B. sup (f b) a)" | 
| 62373 | 225 | by (simp add: sup_ccINF sup_commute) | 
| 226 | ||
| 227 | lemma ccSUP_inf: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 228 | "countable B \<Longrightarrow> inf (SUP b:B. f b) a = (SUP b:B. inf (f b) a)" | 
| 62373 | 229 | by (simp add: inf_ccSUP inf_commute) | 
| 230 | ||
| 231 | lemma ccINF_sup_distrib2: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 232 | "countable A \<Longrightarrow> countable B \<Longrightarrow> sup (INF a:A. f a) (INF b:B. g b) = (INF a:A. INF b:B. sup (f a) (g b))" | 
| 62373 | 233 | by (subst ccINF_commute) (simp_all add: sup_ccINF ccINF_sup) | 
| 234 | ||
| 235 | lemma ccSUP_inf_distrib2: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 236 | "countable A \<Longrightarrow> countable B \<Longrightarrow> inf (SUP a:A. f a) (SUP b:B. g b) = (SUP a:A. SUP b:B. inf (f a) (g b))" | 
| 62373 | 237 | by (subst ccSUP_commute) (simp_all add: inf_ccSUP ccSUP_inf) | 
| 238 | ||
| 239 | context | |
| 240 | fixes f :: "'a \<Rightarrow> 'b::countable_complete_lattice" | |
| 241 | assumes "mono f" | |
| 242 | begin | |
| 243 | ||
| 244 | lemma mono_ccInf: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 245 | "countable A \<Longrightarrow> f (Inf A) \<le> (INF x:A. f x)" | 
| 62373 | 246 | using \<open>mono f\<close> | 
| 247 | by (auto intro!: countable_complete_lattice_class.ccINF_greatest intro: ccInf_lower dest: monoD) | |
| 248 | ||
| 249 | lemma mono_ccSup: | |
| 62374 
cb27a55d868a
remove lattice syntax from countable complete lattice
 hoelzl parents: 
62373diff
changeset | 250 | "countable A \<Longrightarrow> (SUP x:A. f x) \<le> f (Sup A)" | 
| 62373 | 251 | using \<open>mono f\<close> by (auto intro: countable_complete_lattice_class.ccSUP_least ccSup_upper dest: monoD) | 
| 252 | ||
| 253 | lemma mono_ccINF: | |
| 254 | "countable I \<Longrightarrow> f (INF i : I. A i) \<le> (INF x : I. f (A x))" | |
| 255 | by (intro countable_complete_lattice_class.ccINF_greatest monoD[OF \<open>mono f\<close>] ccINF_lower) | |
| 256 | ||
| 257 | lemma mono_ccSUP: | |
| 258 | "countable I \<Longrightarrow> (SUP x : I. f (A x)) \<le> f (SUP i : I. A i)" | |
| 259 | by (intro countable_complete_lattice_class.ccSUP_least monoD[OF \<open>mono f\<close>] ccSUP_upper) | |
| 260 | ||
| 261 | end | |
| 262 | ||
| 263 | end | |
| 264 | ||
| 265 | subsubsection \<open>Instances of countable complete lattices\<close> | |
| 266 | ||
| 267 | instance "fun" :: (type, countable_complete_lattice) countable_complete_lattice | |
| 268 | by standard | |
| 269 | (auto simp: le_fun_def intro!: ccSUP_upper ccSUP_least ccINF_lower ccINF_greatest) | |
| 270 | ||
| 271 | subclass (in complete_lattice) countable_complete_lattice | |
| 272 | by standard (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest) | |
| 273 | ||
| 274 | subclass (in complete_distrib_lattice) countable_complete_distrib_lattice | |
| 275 | by standard (auto intro: sup_Inf inf_Sup) | |
| 276 | ||
| 62390 | 277 | end |