--- 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