# HG changeset patch # User wenzelm # Date 1491335154 -7200 # Node ID 76a96e32bd237a1ba82977fc4ee72bcacdc560fc # Parent 4bb51e6334ede469350ab3c041f91f975d2f2647 tuned (see also 1fa1023b13b9); diff -r 4bb51e6334ed -r 76a96e32bd23 src/HOL/HOLCF/Completion.thy --- 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: "\\u. S <<| u; S <| x\ \ lub S \ x" by (erule exE, drule is_lub_lub, erule is_lubD2) + subsection \Locale for ideal completion\ +hide_const (open) Filter.principal + locale ideal_completion = preorder + fixes principal :: "'a::type \ 'b::cpo" fixes rep :: "'b::cpo \ 'a::type set" diff -r 4bb51e6334ed -r 76a96e32bd23 src/HOL/HOLCF/Plain_HOLCF.thy --- 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. \ -hide_const (open) Filter.principal - end