diff -r 7638776e0977 -r 2cc651d3ce8e src/Doc/Tutorial/Documents/Documents.thy --- a/src/Doc/Tutorial/Documents/Documents.thy Fri Sep 27 22:44:30 2024 +0200 +++ b/src/Doc/Tutorial/Documents/Documents.thy Fri Sep 27 23:47:45 2024 +0200 @@ -37,7 +37,7 @@ operation on boolean values illustrates typical infix declarations. \ -definition xor :: "bool \ bool \ bool" (infixl \[+]\ 60) +definition xor :: "bool \ bool \ bool" (infixl "[+]" 60) where "A [+] B \ (A \ \ B) \ (\ A \ B)" text \ @@ -138,7 +138,7 @@ hide_const xor setup \Sign.add_path "version1"\ (*>*) -definition xor :: "bool \ bool \ bool" (infixl \\\ 60) +definition xor :: "bool \ bool \ bool" (infixl "\" 60) where "A \ B \ (A \ \ B) \ (\ A \ B)" (*<*) setup \Sign.local_path\ @@ -156,10 +156,10 @@ hide_const xor setup \Sign.add_path "version2"\ (*>*) -definition xor :: "bool \ bool \ bool" (infixl \[+]\\ 60) +definition xor :: "bool \ bool \ bool" (infixl "[+]\" 60) where "A [+]\ B \ (A \ \ B) \ (\ A \ B)" -notation (xsymbols) xor (infixl \\\\ 60) +notation (xsymbols) xor (infixl "\\" 60) (*<*) setup \Sign.local_path\ (*>*) @@ -186,10 +186,10 @@ \ datatype currency = - Euro nat (\\\) - | Pounds nat (\\\) - | Yen nat (\\\) - | Dollar nat (\$\) + Euro nat ("\") + | Pounds nat ("\") + | Yen nat ("\") + | Dollar nat ("$") text \ \noindent Here the mixfix annotations on the rightmost column happen @@ -223,7 +223,7 @@ \x \ y\. We assume that a constant \sim\ of type \<^typ>\('a \ 'a) set\ has been introduced at this point.\ (*<*)consts sim :: "('a \ 'a) set"(*>*) -abbreviation sim2 :: "'a \ 'a \ bool" (infix \\\ 50) +abbreviation sim2 :: "'a \ 'a \ bool" (infix "\" 50) where "x \ y \ (x, y) \ sim" text \\noindent The given meta-equality is used as a rewrite rule @@ -238,10 +238,10 @@ stems from Isabelle/HOL itself: \ -abbreviation not_equal :: "'a \ 'a \ bool" (infixl \~=\\ 50) +abbreviation not_equal :: "'a \ 'a \ bool" (infixl "~=\" 50) where "x ~=\ y \ \ (x = y)" -notation (xsymbols) not_equal (infix \\\\ 50) +notation (xsymbols) not_equal (infix "\\" 50) text \\noindent The notation \\\ is introduced separately to restrict it to the \emph{xsymbols} mode.