bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax;
--- a/src/Doc/Main/Main_Doc.thy Mon Sep 20 23:15:02 2021 +0200
+++ b/src/Doc/Main/Main_Doc.thy Tue Sep 21 00:20:47 2021 +0200
@@ -89,7 +89,7 @@
\subsubsection*{Syntax}
-Available by loading theory \<open>Lattice_Syntax\<close> in directory \<open>Library\<close>.
+Available via \<^theory_text>\<open>unbundle lattice_syntax\<close>.
\begin{supertabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}}
@{text[source]"x \<sqsubseteq> y"} & \<^term>\<open>x \<le> y\<close>\\
--- a/src/Doc/Typeclass_Hierarchy/Setup.thy Mon Sep 20 23:15:02 2021 +0200
+++ b/src/Doc/Typeclass_Hierarchy/Setup.thy Tue Sep 21 00:20:47 2021 +0200
@@ -1,7 +1,9 @@
theory Setup
-imports Complex_Main "HOL-Library.Multiset" "HOL-Library.Lattice_Syntax"
+imports Complex_Main "HOL-Library.Multiset"
begin
+unbundle lattice_syntax
+
ML_file \<open>../antiquote_setup.ML\<close>
ML_file \<open>../more_antiquote.ML\<close>
--- a/src/HOL/Examples/Knaster_Tarski.thy Mon Sep 20 23:15:02 2021 +0200
+++ b/src/HOL/Examples/Knaster_Tarski.thy Tue Sep 21 00:20:47 2021 +0200
@@ -7,9 +7,11 @@
section \<open>Textbook-style reasoning: the Knaster-Tarski Theorem\<close>
theory Knaster_Tarski
- imports Main "HOL-Library.Lattice_Syntax"
+ imports Main
begin
+unbundle lattice_syntax
+
subsection \<open>Prose version\<close>
--- a/src/HOL/Library/Combine_PER.thy Mon Sep 20 23:15:02 2021 +0200
+++ b/src/HOL/Library/Combine_PER.thy Tue Sep 21 00:20:47 2021 +0200
@@ -3,9 +3,11 @@
section \<open>A combinator to build partial equivalence relations from a predicate and an equivalence relation\<close>
theory Combine_PER
- imports Main Lattice_Syntax
+ imports Main
begin
+unbundle lattice_syntax
+
definition combine_per :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
where "combine_per P R = (\<lambda>x y. P x \<and> P y) \<sqinter> R"
--- a/src/HOL/Library/Complete_Partial_Order2.thy Mon Sep 20 23:15:02 2021 +0200
+++ b/src/HOL/Library/Complete_Partial_Order2.thy Tue Sep 21 00:20:47 2021 +0200
@@ -5,9 +5,11 @@
section \<open>Formalisation of chain-complete partial orders, continuity and admissibility\<close>
theory Complete_Partial_Order2 imports
- Main Lattice_Syntax
+ Main
begin
+unbundle lattice_syntax
+
lemma chain_transfer [transfer_rule]:
includes lifting_syntax
shows "((A ===> A ===> (=)) ===> rel_set A ===> (=)) Complete_Partial_Order.chain Complete_Partial_Order.chain"
--- a/src/HOL/Library/Lattice_Syntax.thy Mon Sep 20 23:15:02 2021 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,23 +0,0 @@
-(* Author: Florian Haftmann, TU Muenchen *)
-
-section \<open>Pretty syntax for lattice operations\<close>
-
-theory Lattice_Syntax
-imports Main
-begin
-
-notation
- bot ("\<bottom>") and
- top ("\<top>") and
- inf (infixl "\<sqinter>" 70) and
- sup (infixl "\<squnion>" 65) and
- Inf ("\<Sqinter> _" [900] 900) and
- Sup ("\<Squnion> _" [900] 900)
-
-syntax
- "_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10)
- "_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
- "_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10)
- "_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
-
-end
--- a/src/HOL/Library/Library.thy Mon Sep 20 23:15:02 2021 +0200
+++ b/src/HOL/Library/Library.thy Tue Sep 21 00:20:47 2021 +0200
@@ -46,7 +46,6 @@
IArray
Landau_Symbols
Lattice_Algebras
- Lattice_Syntax
Lattice_Constructions
Linear_Temporal_Logic_on_Streams
ListVector
--- a/src/HOL/Library/Option_ord.thy Mon Sep 20 23:15:02 2021 +0200
+++ b/src/HOL/Library/Option_ord.thy Tue Sep 21 00:20:47 2021 +0200
@@ -8,20 +8,7 @@
imports Main
begin
-notation
- bot ("\<bottom>") and
- top ("\<top>") and
- inf (infixl "\<sqinter>" 70) and
- sup (infixl "\<squnion>" 65) and
- Inf ("\<Sqinter> _" [900] 900) and
- Sup ("\<Squnion> _" [900] 900)
-
-syntax
- "_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10)
- "_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
- "_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10)
- "_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
-
+unbundle lattice_syntax
instantiation option :: (preorder) preorder
begin
@@ -462,19 +449,6 @@
instance option :: (complete_linorder) complete_linorder ..
-
-no_notation
- bot ("\<bottom>") and
- top ("\<top>") and
- inf (infixl "\<sqinter>" 70) and
- sup (infixl "\<squnion>" 65) and
- Inf ("\<Sqinter> _" [900] 900) and
- Sup ("\<Squnion> _" [900] 900)
-
-no_syntax
- "_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10)
- "_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
- "_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10)
- "_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
+unbundle no_lattice_syntax
end
--- a/src/HOL/Main.thy Mon Sep 20 23:15:02 2021 +0200
+++ b/src/HOL/Main.thy Tue Sep 21 00:20:47 2021 +0200
@@ -19,7 +19,7 @@
GCD
begin
-text \<open>Namespace cleanup\<close>
+subsection \<open>Namespace cleanup\<close>
hide_const (open)
czero cinfinite cfinite csum cone ctwo Csum cprod cexp image2 image2p vimage2p Gr Grp collect
@@ -29,15 +29,9 @@
hide_fact (open) id_bnf_def type_definition_id_bnf_UNIV
-text \<open>Syntax cleanup\<close>
+subsection \<open>Syntax cleanup\<close>
no_notation
- bot ("\<bottom>") and
- top ("\<top>") and
- inf (infixl "\<sqinter>" 70) and
- sup (infixl "\<squnion>" 65) and
- Inf ("\<Sqinter>") and
- Sup ("\<Squnion>") and
ordLeq2 (infix "<=o" 50) and
ordLeq3 (infix "\<le>o" 50) and
ordLess2 (infix "<o" 50) and
@@ -48,7 +42,9 @@
BNF_Cardinal_Arithmetic.cexp (infixr "^c" 90) and
BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
-bundle cardinal_syntax begin
+bundle cardinal_syntax
+begin
+
notation
ordLeq2 (infix "<=o" 50) and
ordLeq3 (infix "\<le>o" 50) and
@@ -63,8 +59,42 @@
alias czero = BNF_Cardinal_Arithmetic.czero
alias cone = BNF_Cardinal_Arithmetic.cone
alias ctwo = BNF_Cardinal_Arithmetic.ctwo
+
end
+
+subsection \<open>Lattice syntax\<close>
+
+bundle lattice_syntax
+begin
+
+notation
+ bot ("\<bottom>") and
+ top ("\<top>") and
+ inf (infixl "\<sqinter>" 70) and
+ sup (infixl "\<squnion>" 65) and
+ Inf ("\<Sqinter> _" [900] 900) and
+ Sup ("\<Squnion> _" [900] 900)
+
+syntax
+ "_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10)
+ "_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
+ "_SUP1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_./ _)" [0, 10] 10)
+ "_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
+
+end
+
+bundle no_lattice_syntax
+begin
+
+no_notation
+ bot ("\<bottom>") and
+ top ("\<top>") and
+ inf (infixl "\<sqinter>" 70) and
+ sup (infixl "\<squnion>" 65) and
+ Inf ("\<Sqinter> _" [900] 900) and
+ Sup ("\<Squnion> _" [900] 900)
+
no_syntax
"_INF1" :: "pttrns \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_./ _)" [0, 10] 10)
"_INF" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sqinter>_\<in>_./ _)" [0, 0, 10] 10)
@@ -72,3 +102,8 @@
"_SUP" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0, 0, 10] 10)
end
+
+unbundle no_lattice_syntax
+no_notation Inf ("\<Sqinter>") and Sup ("\<Squnion>")
+
+end