diff -r 01b8c8d17f13 -r 2a77bc3b4eac src/FOL/IFOL.thy --- a/src/FOL/IFOL.thy Fri Sep 20 23:36:33 2024 +0200 +++ b/src/FOL/IFOL.thy Fri Sep 20 23:37:00 2024 +0200 @@ -101,7 +101,7 @@ abbreviation not_equal :: \['a, 'a] \ o\ (infixl \\\ 50) where \x \ y \ \ (x = y)\ -syntax "_Uniq" :: "pttrn \ o \ o" (\(2\\<^sub>\\<^sub>1 _./ _)\ [0, 10] 10) +syntax "_Uniq" :: "pttrn \ o \ o" (\(\indent=2 notation=\binder \\<^sub>\\<^sub>1\\\\<^sub>\\<^sub>1 _./ _)\ [0, 10] 10) syntax_consts "_Uniq" \ Uniq translations "\\<^sub>\\<^sub>1x. P" \ "CONST Uniq (\x. P)" @@ -743,10 +743,10 @@ where \Let(s, f) \ f(s)\ syntax - "_bind" :: \[pttrn, 'a] => letbind\ (\(2_ =/ _)\ 10) + "_bind" :: \[pttrn, 'a] => letbind\ (\(\indent=2 notation=\infix letbind\\_ =/ _)\ 10) "" :: \letbind => letbinds\ (\_\) "_binds" :: \[letbind, letbinds] => letbinds\ (\_;/ _\) - "_Let" :: \[letbinds, 'a] => 'a\ (\(let (_)/ in (_))\ 10) + "_Let" :: \[letbinds, 'a] => 'a\ (\(\notation=\mixfix let in\\let (_)/ in (_))\ 10) syntax_consts "_Let" \ Let translations "_Let(_binds(b, bs), e)" == "_Let(b, _Let(bs, e))"