src/HOL/MicroJava/DFA/Listn.thy
changeset 80914 d97fdabd9e2b
parent 68646 7dc9fe795dae
--- a/src/HOL/MicroJava/DFA/Listn.thy	Fri Sep 20 19:07:10 2024 +0200
+++ b/src/HOL/MicroJava/DFA/Listn.thy	Fri Sep 20 19:51:08 2024 +0200
@@ -17,12 +17,12 @@
 
 abbreviation
   lesublist_syntax :: "'a list \<Rightarrow> 'a ord \<Rightarrow> 'a list \<Rightarrow> bool"
-       ("(_ /<=[_] _)" [50, 0, 51] 50)
+       (\<open>(_ /<=[_] _)\<close> [50, 0, 51] 50)
   where "x <=[r] y == x <=_(le r) y"
 
 abbreviation
   lesssublist_syntax :: "'a list \<Rightarrow> 'a ord \<Rightarrow> 'a list \<Rightarrow> bool"
-       ("(_ /<[_] _)" [50, 0, 51] 50)
+       (\<open>(_ /<[_] _)\<close> [50, 0, 51] 50)
   where "x <[r] y == x <_(le r) y"
 
 definition map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list" where
@@ -30,7 +30,7 @@
 
 abbreviation
   plussublist_syntax :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b list \<Rightarrow> 'c list"
-       ("(_ /+[_] _)" [65, 0, 66] 65)
+       (\<open>(_ /+[_] _)\<close> [65, 0, 66] 65)
   where "x +[f] y == x +_(map2 f) y"
 
 primrec coalesce :: "'a err list \<Rightarrow> 'a list err" where