# HG changeset patch # User wenzelm # Date 1632176447 -7200 # Node ID ead56ad40e15b626186ce83a455860597d0ea414 # Parent a9b20bc32fa6b23773c7741bc30d2d63a51aaa29 bundle lattice_syntax / no_lattice_syntax supersedes theory HOL-Library.Lattice_Syntax; diff -r a9b20bc32fa6 -r ead56ad40e15 src/Doc/Main/Main_Doc.thy --- 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 \Lattice_Syntax\ in directory \Library\. +Available via \<^theory_text>\unbundle lattice_syntax\. \begin{supertabular}{@ {} l @ {\quad$\equiv$\quad} l @ {}} @{text[source]"x \ y"} & \<^term>\x \ y\\\ diff -r a9b20bc32fa6 -r ead56ad40e15 src/Doc/Typeclass_Hierarchy/Setup.thy --- 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 \../antiquote_setup.ML\ ML_file \../more_antiquote.ML\ diff -r a9b20bc32fa6 -r ead56ad40e15 src/HOL/Examples/Knaster_Tarski.thy --- 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 \Textbook-style reasoning: the Knaster-Tarski Theorem\ theory Knaster_Tarski - imports Main "HOL-Library.Lattice_Syntax" + imports Main begin +unbundle lattice_syntax + subsection \Prose version\ diff -r a9b20bc32fa6 -r ead56ad40e15 src/HOL/Library/Combine_PER.thy --- 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 \A combinator to build partial equivalence relations from a predicate and an equivalence relation\ theory Combine_PER - imports Main Lattice_Syntax + imports Main begin +unbundle lattice_syntax + definition combine_per :: "('a \ bool) \ ('a \ 'a \ bool) \ 'a \ 'a \ bool" where "combine_per P R = (\x y. P x \ P y) \ R" diff -r a9b20bc32fa6 -r ead56ad40e15 src/HOL/Library/Complete_Partial_Order2.thy --- 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 \Formalisation of chain-complete partial orders, continuity and admissibility\ 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" diff -r a9b20bc32fa6 -r ead56ad40e15 src/HOL/Library/Lattice_Syntax.thy --- 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 \Pretty syntax for lattice operations\ - -theory Lattice_Syntax -imports Main -begin - -notation - bot ("\") and - top ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\ _" [900] 900) and - Sup ("\ _" [900] 900) - -syntax - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - -end diff -r a9b20bc32fa6 -r ead56ad40e15 src/HOL/Library/Library.thy --- 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 diff -r a9b20bc32fa6 -r ead56ad40e15 src/HOL/Library/Option_ord.thy --- 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 ("\") and - top ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\ _" [900] 900) and - Sup ("\ _" [900] 900) - -syntax - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [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 ("\") and - top ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\ _" [900] 900) and - Sup ("\ _" [900] 900) - -no_syntax - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) +unbundle no_lattice_syntax end diff -r a9b20bc32fa6 -r ead56ad40e15 src/HOL/Main.thy --- 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 \Namespace cleanup\ +subsection \Namespace cleanup\ 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 \Syntax cleanup\ +subsection \Syntax cleanup\ no_notation - bot ("\") and - top ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\") and - Sup ("\") and ordLeq2 (infix "<=o" 50) and ordLeq3 (infix "\o" 50) and ordLess2 (infix "(_,/ _)\") -bundle cardinal_syntax begin +bundle cardinal_syntax +begin + notation ordLeq2 (infix "<=o" 50) and ordLeq3 (infix "\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 \Lattice syntax\ + +bundle lattice_syntax +begin + +notation + bot ("\") and + top ("\") and + inf (infixl "\" 70) and + sup (infixl "\" 65) and + Inf ("\ _" [900] 900) and + Sup ("\ _" [900] 900) + +syntax + "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + +end + +bundle no_lattice_syntax +begin + +no_notation + bot ("\") and + top ("\") and + inf (infixl "\" 70) and + sup (infixl "\" 65) and + Inf ("\ _" [900] 900) and + Sup ("\ _" [900] 900) + no_syntax "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) @@ -72,3 +102,8 @@ "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) end + +unbundle no_lattice_syntax +no_notation Inf ("\") and Sup ("\") + +end