--- a/src/HOL/HOLCF/Completion.thy Tue Apr 04 21:44:44 2017 +0200
+++ b/src/HOL/HOLCF/Completion.thy Tue Apr 04 21:45:54 2017 +0200
@@ -137,8 +137,11 @@
lemma is_lub_thelub_ex: "\<lbrakk>\<exists>u. S <<| u; S <| x\<rbrakk> \<Longrightarrow> lub S \<sqsubseteq> x"
by (erule exE, drule is_lub_lub, erule is_lubD2)
+
subsection \<open>Locale for ideal completion\<close>
+hide_const (open) Filter.principal
+
locale ideal_completion = preorder +
fixes principal :: "'a::type \<Rightarrow> 'b::cpo"
fixes rep :: "'b::cpo \<Rightarrow> 'a::type set"
--- a/src/HOL/HOLCF/Plain_HOLCF.thy Tue Apr 04 21:44:44 2017 +0200
+++ b/src/HOL/HOLCF/Plain_HOLCF.thy Tue Apr 04 21:45:54 2017 +0200
@@ -12,6 +12,4 @@
Basic HOLCF concepts and types; does not include definition packages.
\<close>
-hide_const (open) Filter.principal
-
end