remove lattice syntax from countable complete lattice
authorhoelzl
Fri, 19 Feb 2016 12:25:57 +0100
changeset 62374 cb27a55d868a
parent 62373 ea7a442e9a56
child 62375 670063003ad3
remove lattice syntax from countable complete lattice
src/HOL/Library/Countable_Complete_Lattices.thy
src/HOL/Library/Extended_Nat.thy
src/HOL/Library/Order_Continuity.thy
--- a/src/HOL/Library/Countable_Complete_Lattices.thy	Thu Feb 18 13:54:44 2016 +0100
+++ b/src/HOL/Library/Countable_Complete_Lattices.thy	Fri Feb 19 12:25:57 2016 +0100
@@ -5,169 +5,169 @@
 section \<open>Countable Complete Lattices\<close>
 
 theory Countable_Complete_Lattices
-  imports Main Lattice_Syntax Countable_Set
+  imports Main Countable_Set
 begin
 
 lemma UNIV_nat_eq: "UNIV = insert 0 (range Suc)"
   by (metis UNIV_eq_I nat.nchotomy insertCI rangeI)
 
 class countable_complete_lattice = lattice + Inf + Sup + bot + top +
-  assumes ccInf_lower: "countable A \<Longrightarrow> x \<in> A \<Longrightarrow> \<Sqinter>A \<le> x"
-  assumes ccInf_greatest: "countable A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> \<Sqinter>A"
-  assumes ccSup_upper: "countable A \<Longrightarrow> x \<in> A \<Longrightarrow> x \<le> \<Squnion>A"
-  assumes ccSup_least: "countable A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> \<Squnion>A \<le> z"
-  assumes ccInf_empty [simp]: "\<Sqinter>{} = \<top>"
-  assumes ccSup_empty [simp]: "\<Squnion>{} = \<bottom>"
+  assumes ccInf_lower: "countable A \<Longrightarrow> x \<in> A \<Longrightarrow> Inf A \<le> x"
+  assumes ccInf_greatest: "countable A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf A"
+  assumes ccSup_upper: "countable A \<Longrightarrow> x \<in> A \<Longrightarrow> x \<le> Sup A"
+  assumes ccSup_least: "countable A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup A \<le> z"
+  assumes ccInf_empty [simp]: "Inf {} = top"
+  assumes ccSup_empty [simp]: "Sup {} = bot"
 begin
 
 subclass bounded_lattice
 proof
   fix a
-  show "\<bottom> \<le> a" by (auto intro: ccSup_least simp only: ccSup_empty [symmetric])
-  show "a \<le> \<top>" by (auto intro: ccInf_greatest simp only: ccInf_empty [symmetric])
+  show "bot \<le> a" by (auto intro: ccSup_least simp only: ccSup_empty [symmetric])
+  show "a \<le> top" by (auto intro: ccInf_greatest simp only: ccInf_empty [symmetric])
 qed
 
-lemma ccINF_lower: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> (\<Sqinter>i\<in>A. f i) \<le> f i"
+lemma ccINF_lower: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> (INF i :A. f i) \<le> f i"
   using ccInf_lower [of "f ` A"] by simp
 
-lemma ccINF_greatest: "countable A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> u \<le> f i) \<Longrightarrow> u \<le> (\<Sqinter>i\<in>A. f i)"
+lemma ccINF_greatest: "countable A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> u \<le> f i) \<Longrightarrow> u \<le> (INF i :A. f i)"
   using ccInf_greatest [of "f ` A"] by auto
 
-lemma ccSUP_upper: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> f i \<le> (\<Squnion>i\<in>A. f i)"
+lemma ccSUP_upper: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> f i \<le> (SUP i :A. f i)"
   using ccSup_upper [of "f ` A"] by simp
 
-lemma ccSUP_least: "countable A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> f i \<le> u) \<Longrightarrow> (\<Squnion>i\<in>A. f i) \<le> u"
+lemma ccSUP_least: "countable A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> f i \<le> u) \<Longrightarrow> (SUP i :A. f i) \<le> u"
   using ccSup_least [of "f ` A"] by auto
 
-lemma ccInf_lower2: "countable A \<Longrightarrow> u \<in> A \<Longrightarrow> u \<le> v \<Longrightarrow> \<Sqinter>A \<le> v"
+lemma ccInf_lower2: "countable A \<Longrightarrow> u \<in> A \<Longrightarrow> u \<le> v \<Longrightarrow> Inf A \<le> v"
   using ccInf_lower [of A u] by auto
 
-lemma ccINF_lower2: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> f i \<le> u \<Longrightarrow> (\<Sqinter>i\<in>A. f i) \<le> u"
+lemma ccINF_lower2: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> f i \<le> u \<Longrightarrow> (INF i :A. f i) \<le> u"
   using ccINF_lower [of A i f] by auto
 
-lemma ccSup_upper2: "countable A \<Longrightarrow> u \<in> A \<Longrightarrow> v \<le> u \<Longrightarrow> v \<le> \<Squnion>A"
+lemma ccSup_upper2: "countable A \<Longrightarrow> u \<in> A \<Longrightarrow> v \<le> u \<Longrightarrow> v \<le> Sup A"
   using ccSup_upper [of A u] by auto
 
-lemma ccSUP_upper2: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> u \<le> f i \<Longrightarrow> u \<le> (\<Squnion>i\<in>A. f i)"
+lemma ccSUP_upper2: "countable A \<Longrightarrow> i \<in> A \<Longrightarrow> u \<le> f i \<Longrightarrow> u \<le> (SUP i :A. f i)"
   using ccSUP_upper [of A i f] by auto
 
-lemma le_ccInf_iff: "countable A \<Longrightarrow> b \<le> \<Sqinter>A \<longleftrightarrow> (\<forall>a\<in>A. b \<le> a)"
+lemma le_ccInf_iff: "countable A \<Longrightarrow> b \<le> Inf A \<longleftrightarrow> (\<forall>a\<in>A. b \<le> a)"
   by (auto intro: ccInf_greatest dest: ccInf_lower)
 
-lemma le_ccINF_iff: "countable A \<Longrightarrow> u \<le> (\<Sqinter>i\<in>A. f i) \<longleftrightarrow> (\<forall>i\<in>A. u \<le> f i)"
+lemma le_ccINF_iff: "countable A \<Longrightarrow> u \<le> (INF i :A. f i) \<longleftrightarrow> (\<forall>i\<in>A. u \<le> f i)"
   using le_ccInf_iff [of "f ` A"] by simp
 
-lemma ccSup_le_iff: "countable A \<Longrightarrow> \<Squnion>A \<le> b \<longleftrightarrow> (\<forall>a\<in>A. a \<le> b)"
+lemma ccSup_le_iff: "countable A \<Longrightarrow> Sup A \<le> b \<longleftrightarrow> (\<forall>a\<in>A. a \<le> b)"
   by (auto intro: ccSup_least dest: ccSup_upper)
 
-lemma ccSUP_le_iff: "countable A \<Longrightarrow> (\<Squnion>i\<in>A. f i) \<le> u \<longleftrightarrow> (\<forall>i\<in>A. f i \<le> u)"
+lemma ccSUP_le_iff: "countable A \<Longrightarrow> (SUP i :A. f i) \<le> u \<longleftrightarrow> (\<forall>i\<in>A. f i \<le> u)"
   using ccSup_le_iff [of "f ` A"] by simp
 
-lemma ccInf_insert [simp]: "countable A \<Longrightarrow> \<Sqinter>insert a A = a \<sqinter> \<Sqinter>A"
+lemma ccInf_insert [simp]: "countable A \<Longrightarrow> Inf (insert a A) = inf a (Inf A)"
   by (force intro: le_infI le_infI1 le_infI2 antisym ccInf_greatest ccInf_lower)
 
-lemma ccINF_insert [simp]: "countable A \<Longrightarrow> (\<Sqinter>x\<in>insert a A. f x) = f a \<sqinter> INFIMUM A f"
+lemma ccINF_insert [simp]: "countable A \<Longrightarrow> (INF x:insert a A. f x) = inf (f a) (INFIMUM A f)"
   unfolding image_insert by simp
 
-lemma ccSup_insert [simp]: "countable A \<Longrightarrow> \<Squnion>insert a A = a \<squnion> \<Squnion>A"
+lemma ccSup_insert [simp]: "countable A \<Longrightarrow> Sup (insert a A) = sup a (Sup A)"
   by (force intro: le_supI le_supI1 le_supI2 antisym ccSup_least ccSup_upper)
 
-lemma ccSUP_insert [simp]: "countable A \<Longrightarrow> (\<Squnion>x\<in>insert a A. f x) = f a \<squnion> SUPREMUM A f"
+lemma ccSUP_insert [simp]: "countable A \<Longrightarrow> (SUP x:insert a A. f x) = sup (f a) (SUPREMUM A f)"
   unfolding image_insert by simp
 
-lemma ccINF_empty [simp]: "(\<Sqinter>x\<in>{}. f x) = \<top>"
+lemma ccINF_empty [simp]: "(INF x:{}. f x) = top"
   unfolding image_empty by simp
 
-lemma ccSUP_empty [simp]: "(\<Squnion>x\<in>{}. f x) = \<bottom>"
+lemma ccSUP_empty [simp]: "(SUP x:{}. f x) = bot"
   unfolding image_empty by simp
 
-lemma ccInf_superset_mono: "countable A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> \<Sqinter>A \<le> \<Sqinter>B"
+lemma ccInf_superset_mono: "countable A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> Inf A \<le> Inf B"
   by (auto intro: ccInf_greatest ccInf_lower countable_subset)
 
-lemma ccSup_subset_mono: "countable B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> \<Squnion>A \<le> \<Squnion>B"
+lemma ccSup_subset_mono: "countable B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
   by (auto intro: ccSup_least ccSup_upper countable_subset)
 
 lemma ccInf_mono:
   assumes [intro]: "countable B" "countable A"
   assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b"
-  shows "\<Sqinter>A \<le> \<Sqinter>B"
+  shows "Inf A \<le> Inf B"
 proof (rule ccInf_greatest)
   fix b assume "b \<in> B"
   with assms obtain a where "a \<in> A" and "a \<le> b" by blast
-  from \<open>a \<in> A\<close> have "\<Sqinter>A \<le> a" by (rule ccInf_lower[rotated]) auto
-  with \<open>a \<le> b\<close> show "\<Sqinter>A \<le> b" by auto
+  from \<open>a \<in> A\<close> have "Inf A \<le> a" by (rule ccInf_lower[rotated]) auto
+  with \<open>a \<le> b\<close> show "Inf A \<le> b" by auto
 qed auto
 
 lemma ccINF_mono:
-  "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> (\<Sqinter>n\<in>A. f n) \<le> (\<Sqinter>n\<in>B. g n)"
+  "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)"
   using ccInf_mono [of "g ` B" "f ` A"] by auto
 
 lemma ccSup_mono:
   assumes [intro]: "countable B" "countable A"
   assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<le> b"
-  shows "\<Squnion>A \<le> \<Squnion>B"
+  shows "Sup A \<le> Sup B"
 proof (rule ccSup_least)
   fix a assume "a \<in> A"
   with assms obtain b where "b \<in> B" and "a \<le> b" by blast
-  from \<open>b \<in> B\<close> have "b \<le> \<Squnion>B" by (rule ccSup_upper[rotated]) auto
-  with \<open>a \<le> b\<close> show "a \<le> \<Squnion>B" by auto
+  from \<open>b \<in> B\<close> have "b \<le> Sup B" by (rule ccSup_upper[rotated]) auto
+  with \<open>a \<le> b\<close> show "a \<le> Sup B" by auto
 qed auto
 
 lemma ccSUP_mono:
-  "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> (\<Squnion>n\<in>A. f n) \<le> (\<Squnion>n\<in>B. g n)"
+  "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)"
   using ccSup_mono [of "g ` B" "f ` A"] by auto
 
 lemma ccINF_superset_mono:
-  "countable A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (\<Sqinter>x\<in>A. f x) \<le> (\<Sqinter>x\<in>B. g x)"
+  "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)"
   by (blast intro: ccINF_mono countable_subset dest: subsetD)
 
 lemma ccSUP_subset_mono:
-  "countable B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (\<Squnion>x\<in>A. f x) \<le> (\<Squnion>x\<in>B. g x)"
+  "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)"
   by (blast intro: ccSUP_mono countable_subset dest: subsetD)
 
 
-lemma less_eq_ccInf_inter: "countable A \<Longrightarrow> countable B \<Longrightarrow> \<Sqinter>A \<squnion> \<Sqinter>B \<le> \<Sqinter>(A \<inter> B)"
+lemma less_eq_ccInf_inter: "countable A \<Longrightarrow> countable B \<Longrightarrow> sup (Inf A) (Inf B) \<le> Inf (A \<inter> B)"
   by (auto intro: ccInf_greatest ccInf_lower)
 
-lemma ccSup_inter_less_eq: "countable A \<Longrightarrow> countable B \<Longrightarrow> \<Squnion>(A \<inter> B) \<le> \<Squnion>A \<sqinter> \<Squnion>B "
+lemma ccSup_inter_less_eq: "countable A \<Longrightarrow> countable B \<Longrightarrow> Sup (A \<inter> B) \<le> inf (Sup A) (Sup B)"
   by (auto intro: ccSup_least ccSup_upper)
 
-lemma ccInf_union_distrib: "countable A \<Longrightarrow> countable B \<Longrightarrow> \<Sqinter>(A \<union> B) = \<Sqinter>A \<sqinter> \<Sqinter>B"
+lemma ccInf_union_distrib: "countable A \<Longrightarrow> countable B \<Longrightarrow> Inf (A \<union> B) = inf (Inf A) (Inf B)"
   by (rule antisym) (auto intro: ccInf_greatest ccInf_lower le_infI1 le_infI2)
 
 lemma ccINF_union:
-  "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<Sqinter>i \<in> A \<union> B. M i) = (\<Sqinter>i \<in> A. M i) \<sqinter> (\<Sqinter>i\<in>B. M i)"
+  "countable A \<Longrightarrow> countable B \<Longrightarrow> (INF i:A \<union> B. M i) = inf (INF i:A. M i) (INF i:B. M i)"
   by (auto intro!: antisym ccINF_mono intro: le_infI1 le_infI2 ccINF_greatest ccINF_lower)
 
-lemma ccSup_union_distrib: "countable A \<Longrightarrow> countable B \<Longrightarrow> \<Squnion>(A \<union> B) = \<Squnion>A \<squnion> \<Squnion>B"
+lemma ccSup_union_distrib: "countable A \<Longrightarrow> countable B \<Longrightarrow> Sup (A \<union> B) = sup (Sup A) (Sup B)"
   by (rule antisym) (auto intro: ccSup_least ccSup_upper le_supI1 le_supI2)
 
 lemma ccSUP_union:
-  "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<Squnion>i \<in> A \<union> B. M i) = (\<Squnion>i \<in> A. M i) \<squnion> (\<Squnion>i\<in>B. M i)"
+  "countable A \<Longrightarrow> countable B \<Longrightarrow> (SUP i:A \<union> B. M i) = sup (SUP i:A. M i) (SUP i:B. M i)"
   by (auto intro!: antisym ccSUP_mono intro: le_supI1 le_supI2 ccSUP_least ccSUP_upper)
 
-lemma ccINF_inf_distrib: "countable A \<Longrightarrow> (\<Sqinter>a\<in>A. f a) \<sqinter> (\<Sqinter>a\<in>A. g a) = (\<Sqinter>a\<in>A. f a \<sqinter> g a)"
+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))"
   by (rule antisym) (rule ccINF_greatest, auto intro: le_infI1 le_infI2 ccINF_lower ccINF_mono)
 
-lemma ccSUP_sup_distrib: "countable A \<Longrightarrow> (\<Squnion>a\<in>A. f a) \<squnion> (\<Squnion>a\<in>A. g a) = (\<Squnion>a\<in>A. f a \<squnion> g a)"
+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))"
   by (rule antisym[rotated]) (rule ccSUP_least, auto intro: le_supI1 le_supI2 ccSUP_upper ccSUP_mono)
 
-lemma ccINF_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Sqinter>i\<in>A. f) = f"
+lemma ccINF_const [simp]: "A \<noteq> {} \<Longrightarrow> (INF i :A. f) = f"
   unfolding image_constant_conv by auto
 
-lemma ccSUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (\<Squnion>i\<in>A. f) = f"
+lemma ccSUP_const [simp]: "A \<noteq> {} \<Longrightarrow> (SUP i :A. f) = f"
   unfolding image_constant_conv by auto
 
-lemma ccINF_top [simp]: "(\<Sqinter>x\<in>A. \<top>) = \<top>"
+lemma ccINF_top [simp]: "(INF x:A. top) = top"
   by (cases "A = {}") simp_all
 
-lemma ccSUP_bot [simp]: "(\<Squnion>x\<in>A. \<bottom>) = \<bottom>"
+lemma ccSUP_bot [simp]: "(SUP x:A. bot) = bot"
   by (cases "A = {}") simp_all
 
-lemma ccINF_commute: "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<Sqinter>i\<in>A. \<Sqinter>j\<in>B. f i j) = (\<Sqinter>j\<in>B. \<Sqinter>i\<in>A. f i j)"
+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)"
   by (iprover intro: ccINF_lower ccINF_greatest order_trans antisym)
 
-lemma ccSUP_commute: "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<Squnion>i\<in>A. \<Squnion>j\<in>B. f i j) = (\<Squnion>j\<in>B. \<Squnion>i\<in>A. f i j)"
+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)"
   by (iprover intro: ccSUP_upper ccSUP_least order_trans antisym)
 
 end
@@ -176,62 +176,64 @@
   fixes a :: "'a::{countable_complete_lattice, linorder}"
 begin
 
-lemma less_ccSup_iff: "countable S \<Longrightarrow> a < \<Squnion>S \<longleftrightarrow> (\<exists>x\<in>S. a < x)"
+lemma less_ccSup_iff: "countable S \<Longrightarrow> a < Sup S \<longleftrightarrow> (\<exists>x\<in>S. a < x)"
   unfolding not_le [symmetric] by (subst ccSup_le_iff) auto
 
-lemma less_ccSUP_iff: "countable A \<Longrightarrow> a < (\<Squnion>i\<in>A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
+lemma less_ccSUP_iff: "countable A \<Longrightarrow> a < (SUP i:A. f i) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
   using less_ccSup_iff [of "f ` A"] by simp
 
-lemma ccInf_less_iff: "countable S \<Longrightarrow> \<Sqinter>S < a \<longleftrightarrow> (\<exists>x\<in>S. x < a)"
+lemma ccInf_less_iff: "countable S \<Longrightarrow> Inf S < a \<longleftrightarrow> (\<exists>x\<in>S. x < a)"
   unfolding not_le [symmetric] by (subst le_ccInf_iff) auto
 
-lemma ccINF_less_iff: "countable A \<Longrightarrow> (\<Sqinter>i\<in>A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
+lemma ccINF_less_iff: "countable A \<Longrightarrow> (INF i:A. f i) < a \<longleftrightarrow> (\<exists>x\<in>A. f x < a)"
   using ccInf_less_iff [of "f ` A"] by simp
 
 end
 
 class countable_complete_distrib_lattice = countable_complete_lattice +
-  assumes sup_ccInf: "countable B \<Longrightarrow> a \<squnion> \<Sqinter>B = (\<Sqinter>b\<in>B. a \<squnion> b)"
-  assumes inf_ccSup: "countable B \<Longrightarrow> a \<sqinter> \<Squnion>B = (\<Squnion>b\<in>B. a \<sqinter> b)"
+  assumes sup_ccInf: "countable B \<Longrightarrow> sup a (Inf B) = (INF b:B. sup a b)"
+  assumes inf_ccSup: "countable B \<Longrightarrow> inf a (Sup B) = (SUP b:B. inf a b)"
 begin
 
 lemma sup_ccINF:
-  "countable B \<Longrightarrow> a \<squnion> (\<Sqinter>b\<in>B. f b) = (\<Sqinter>b\<in>B. a \<squnion> f b)"
+  "countable B \<Longrightarrow> sup a (INF b:B. f b) = (INF b:B. sup a (f b))"
   by (simp only: sup_ccInf image_image countable_image)
 
 lemma inf_ccSUP:
-  "countable B \<Longrightarrow> a \<sqinter> (\<Squnion>b\<in>B. f b) = (\<Squnion>b\<in>B. a \<sqinter> f b)"
+  "countable B \<Longrightarrow> inf a (SUP b:B. f b) = (SUP b:B. inf a (f b))"
   by (simp only: inf_ccSup image_image countable_image)
 
-subclass distrib_lattice proof
+subclass distrib_lattice
+proof
   fix a b c
-  from sup_ccInf[of "{b, c}" a] have "a \<squnion> \<Sqinter>{b, c} = (\<Sqinter>d\<in>{b, c}. a \<squnion> d)"
+  from sup_ccInf[of "{b, c}" a] have "sup a (Inf {b, c}) = (INF d:{b, c}. sup a d)"
     by simp
-  then show "a \<squnion> b \<sqinter> c = (a \<squnion> b) \<sqinter> (a \<squnion> c)" by simp
+  then show "sup a (inf b c) = inf (sup a b) (sup a c)"
+    by simp
 qed
 
 lemma ccInf_sup:
-  "countable B \<Longrightarrow> \<Sqinter>B \<squnion> a = (\<Sqinter>b\<in>B. b \<squnion> a)"
+  "countable B \<Longrightarrow> sup (Inf B) a = (INF b:B. sup b a)"
   by (simp add: sup_ccInf sup_commute)
 
 lemma ccSup_inf:
-  "countable B \<Longrightarrow> \<Squnion>B \<sqinter> a = (\<Squnion>b\<in>B. b \<sqinter> a)"
+  "countable B \<Longrightarrow> inf (Sup B) a = (SUP b:B. inf b a)"
   by (simp add: inf_ccSup inf_commute)
 
 lemma ccINF_sup:
-  "countable B \<Longrightarrow> (\<Sqinter>b\<in>B. f b) \<squnion> a = (\<Sqinter>b\<in>B. f b \<squnion> a)"
+  "countable B \<Longrightarrow> sup (INF b:B. f b) a = (INF b:B. sup (f b) a)"
   by (simp add: sup_ccINF sup_commute)
 
 lemma ccSUP_inf:
-  "countable B \<Longrightarrow> (\<Squnion>b\<in>B. f b) \<sqinter> a = (\<Squnion>b\<in>B. f b \<sqinter> a)"
+  "countable B \<Longrightarrow> inf (SUP b:B. f b) a = (SUP b:B. inf (f b) a)"
   by (simp add: inf_ccSUP inf_commute)
 
 lemma ccINF_sup_distrib2:
-  "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<Sqinter>a\<in>A. f a) \<squnion> (\<Sqinter>b\<in>B. g b) = (\<Sqinter>a\<in>A. \<Sqinter>b\<in>B. f a \<squnion> g b)"
+  "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))"
   by (subst ccINF_commute) (simp_all add: sup_ccINF ccINF_sup)
 
 lemma ccSUP_inf_distrib2:
-  "countable A \<Longrightarrow> countable B \<Longrightarrow> (\<Squnion>a\<in>A. f a) \<sqinter> (\<Squnion>b\<in>B. g b) = (\<Squnion>a\<in>A. \<Squnion>b\<in>B. f a \<sqinter> g b)"
+  "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))"
   by (subst ccSUP_commute) (simp_all add: inf_ccSUP ccSUP_inf)
 
 context
@@ -240,12 +242,12 @@
 begin
 
 lemma mono_ccInf:
-  "countable A \<Longrightarrow> f (\<Sqinter>A) \<le> (\<Sqinter>x\<in>A. f x)"
+  "countable A \<Longrightarrow> f (Inf A) \<le> (INF x:A. f x)"
   using \<open>mono f\<close>
   by (auto intro!: countable_complete_lattice_class.ccINF_greatest intro: ccInf_lower dest: monoD)
 
 lemma mono_ccSup:
-  "countable A \<Longrightarrow> (\<Squnion>x\<in>A. f x) \<le> f (\<Squnion>A)"
+  "countable A \<Longrightarrow> (SUP x:A. f x) \<le> f (Sup A)"
   using \<open>mono f\<close> by (auto intro: countable_complete_lattice_class.ccSUP_least ccSup_upper dest: monoD)
 
 lemma mono_ccINF:
--- a/src/HOL/Library/Extended_Nat.thy	Thu Feb 18 13:54:44 2016 +0100
+++ b/src/HOL/Library/Extended_Nat.thy	Fri Feb 19 12:25:57 2016 +0100
@@ -12,7 +12,6 @@
 class infinity =
   fixes infinity :: "'a"  ("\<infinity>")
 
-
 subsection \<open>Type definition\<close>
 
 text \<open>
@@ -26,7 +25,7 @@
 
 definition enat :: "nat \<Rightarrow> enat" where
   "enat n = Abs_enat (Some n)"
- 
+
 instantiation enat :: infinity
 begin
 
@@ -40,7 +39,7 @@
   show "\<exists>to_nat::enat \<Rightarrow> nat. inj to_nat"
     by (rule exI[of _ "to_nat \<circ> Rep_enat"]) (simp add: inj_on_def Rep_enat_inject)
 qed
- 
+
 old_rep_datatype enat "\<infinity> :: enat"
 proof -
   fix P i assume "\<And>j. P (enat j)" "P \<infinity>"
@@ -169,7 +168,7 @@
 lemma eSuc_plus_1:
   "eSuc n = n + 1"
   by (cases n) (simp_all add: eSuc_enat one_enat_def)
-  
+
 lemma plus_1_eSuc:
   "1 + q = eSuc q"
   "q + 1 = eSuc q"
@@ -399,7 +398,7 @@
 
 lemma eSuc_ile_mono [simp]: "eSuc n \<le> eSuc m \<longleftrightarrow> n \<le> m"
   by (simp add: eSuc_def less_eq_enat_def split: enat.splits)
- 
+
 lemma eSuc_mono [simp]: "eSuc n < eSuc m \<longleftrightarrow> n < m"
   by (simp add: eSuc_def less_enat_def split: enat.splits)
 
@@ -479,7 +478,7 @@
 lemma eSuc_max: "eSuc (max x y) = max (eSuc x) (eSuc y)"
   by (simp add: eSuc_def split: enat.split)
 
-lemma eSuc_Max: 
+lemma eSuc_Max:
   assumes "finite A" "A \<noteq> {}"
   shows "eSuc (Max A) = Max (eSuc ` A)"
 using assms proof induction
--- a/src/HOL/Library/Order_Continuity.thy	Thu Feb 18 13:54:44 2016 +0100
+++ b/src/HOL/Library/Order_Continuity.thy	Fri Feb 19 12:25:57 2016 +0100
@@ -393,12 +393,12 @@
 subsubsection \<open>Least fixed points in countable complete lattices\<close>
 
 definition (in countable_complete_lattice) cclfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
-  where "cclfp f = (\<Squnion>i. (f ^^ i) \<bottom>)"
+  where "cclfp f = (SUP i. (f ^^ i) bot)"
 
 lemma cclfp_unfold:
   assumes "sup_continuous F" shows "cclfp F = F (cclfp F)"
 proof -
-  have "cclfp F = (\<Squnion>i. F ((F ^^ i) \<bottom>))"
+  have "cclfp F = (SUP i. F ((F ^^ i) bot))"
     unfolding cclfp_def by (subst UNIV_nat_eq) auto
   also have "\<dots> = F (cclfp F)"
     unfolding cclfp_def
@@ -409,7 +409,7 @@
 lemma cclfp_lowerbound: assumes f: "mono f" and A: "f A \<le> A" shows "cclfp f \<le> A"
   unfolding cclfp_def
 proof (intro ccSUP_least)
-  fix i show "(f ^^ i) \<bottom> \<le> A"
+  fix i show "(f ^^ i) bot \<le> A"
   proof (induction i)
     case (Suc i) from monoD[OF f this] A show ?case
       by auto
@@ -418,12 +418,12 @@
 
 lemma cclfp_transfer:
   assumes "sup_continuous \<alpha>" "mono f"
-  assumes "\<alpha> \<bottom> = \<bottom>" "\<And>x. \<alpha> (f x) = g (\<alpha> x)"
+  assumes "\<alpha> bot = bot" "\<And>x. \<alpha> (f x) = g (\<alpha> x)"
   shows "\<alpha> (cclfp f) = cclfp g"
 proof -
-  have "\<alpha> (cclfp f) = (\<Squnion>i. \<alpha> ((f ^^ i) \<bottom>))"
+  have "\<alpha> (cclfp f) = (SUP i. \<alpha> ((f ^^ i) bot))"
     unfolding cclfp_def by (intro sup_continuousD assms mono_funpow sup_continuous_mono)
-  moreover have "\<alpha> ((f ^^ i) \<bottom>) = (g ^^ i) \<bottom>" for i
+  moreover have "\<alpha> ((f ^^ i) bot) = (g ^^ i) bot" for i
     by (induction i) (simp_all add: assms)
   ultimately show ?thesis
     by (simp add: cclfp_def)