# HG changeset patch # User wenzelm # Date 1726854668 -7200 # Node ID d97fdabd9e2b990954ece27589f7952851ff5c04 # Parent 46f59511b7bbd4b9f4f28620acccd451525634a3 standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing; diff -r 46f59511b7bb -r d97fdabd9e2b src/CCL/CCL.thy --- a/src/CCL/CCL.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CCL/CCL.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,18 +27,18 @@ consts (*** Evaluation Judgement ***) - Eval :: "[i,i]\prop" (infixl "\" 20) + Eval :: "[i,i]\prop" (infixl \\\ 20) (*** Bisimulations for pre-order and equality ***) - po :: "['a,'a]\o" (infixl "[=" 50) + po :: "['a,'a]\o" (infixl \[=\ 50) (*** Term Formers ***) true :: "i" false :: "i" - pair :: "[i,i]\i" ("(1<_,/_>)") - lambda :: "(i\i)\i" (binder "lam " 55) + pair :: "[i,i]\i" (\(1<_,/_>)\) + lambda :: "(i\i)\i" (binder \lam \ 55) "case" :: "[i,i,i,[i,i]\i,(i\i)\i]\i" - "apply" :: "[i,i]\i" (infixl "`" 56) + "apply" :: "[i,i]\i" (infixl \`\ 56) bot :: "i" (******* EVALUATION SEMANTICS *******) diff -r 46f59511b7bb -r d97fdabd9e2b src/CCL/Set.thy --- a/src/CCL/Set.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CCL/Set.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,12 +13,12 @@ subsection \Set comprehension and membership\ axiomatization Collect :: "['a \ o] \ 'a set" - and mem :: "['a, 'a set] \ o" (infixl ":" 50) + and mem :: "['a, 'a set] \ o" (infixl \:\ 50) where mem_Collect_iff: "(a : Collect(P)) \ P(a)" and set_extension: "A = B \ (ALL x. x:A \ x:B)" syntax - "_Coll" :: "[idt, o] \ 'a set" ("(1{_./ _})") + "_Coll" :: "[idt, o] \ 'a set" (\(1{_./ _})\) syntax_consts "_Coll" == Collect translations @@ -50,8 +50,8 @@ where "Bex(A, P) == EX x. x:A \ P(x)" syntax - "_Ball" :: "[idt, 'a set, o] \ o" ("(ALL _:_./ _)" [0, 0, 0] 10) - "_Bex" :: "[idt, 'a set, o] \ o" ("(EX _:_./ _)" [0, 0, 0] 10) + "_Ball" :: "[idt, 'a set, o] \ o" (\(ALL _:_./ _)\ [0, 0, 0] 10) + "_Bex" :: "[idt, 'a set, o] \ o" (\(EX _:_./ _)\ [0, 0, 0] 10) syntax_consts "_Ball" == Ball and "_Bex" == Bex @@ -96,22 +96,22 @@ subsection \Further operations\ -definition subset :: "['a set, 'a set] \ o" (infixl "<=" 50) +definition subset :: "['a set, 'a set] \ o" (infixl \<=\ 50) where "A <= B == ALL x:A. x:B" definition mono :: "['a set \ 'b set] \ o" where "mono(f) == (ALL A B. A <= B \ f(A) <= f(B))" -definition singleton :: "'a \ 'a set" ("{_}") +definition singleton :: "'a \ 'a set" (\{_}\) where "{a} == {x. x=a}" -definition empty :: "'a set" ("{}") +definition empty :: "'a set" (\{}\) where "{} == {x. False}" -definition Un :: "['a set, 'a set] \ 'a set" (infixl "Un" 65) +definition Un :: "['a set, 'a set] \ 'a set" (infixl \Un\ 65) where "A Un B == {x. x:A | x:B}" -definition Int :: "['a set, 'a set] \ 'a set" (infixl "Int" 70) +definition Int :: "['a set, 'a set] \ 'a set" (infixl \Int\ 70) where "A Int B == {x. x:A \ x:B}" definition Compl :: "('a set) \ 'a set" @@ -127,8 +127,8 @@ where "UNION(A, B) == {y. EX x:A. y: B(x)}" syntax - "_INTER" :: "[idt, 'a set, 'b set] \ 'b set" ("(INT _:_./ _)" [0, 0, 0] 10) - "_UNION" :: "[idt, 'a set, 'b set] \ 'b set" ("(UN _:_./ _)" [0, 0, 0] 10) + "_INTER" :: "[idt, 'a set, 'b set] \ 'b set" (\(INT _:_./ _)\ [0, 0, 0] 10) + "_UNION" :: "[idt, 'a set, 'b set] \ 'b set" (\(UN _:_./ _)\ [0, 0, 0] 10) syntax_consts "_INTER" == INTER and "_UNION" == UNION diff -r 46f59511b7bb -r d97fdabd9e2b src/CCL/Term.thy --- a/src/CCL/Term.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CCL/Term.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ definition one :: "i" where "one == true" -definition "if" :: "[i,i,i]\i" ("(3if _/ then _/ else _)" [0,0,60] 60) +definition "if" :: "[i,i,i]\i" (\(3if _/ then _/ else _)\ [0,0,60] 60) where "if b then t else u == case(b, t, u, \ x y. bot, \v. bot)" definition inl :: "i\i" @@ -48,7 +48,7 @@ definition "let1" :: "[i,i\i]\i" where let_def: "let1(t, f) == case(t,f(true),f(false), \x y. f(), \u. f(lam x. u(x)))" -syntax "_let1" :: "[idt,i,i]\i" ("(3let _ be _/ in _)" [0,0,60] 60) +syntax "_let1" :: "[idt,i,i]\i" (\(3let _ be _/ in _)\ [0,0,60] 60) syntax_consts "_let1" == let1 translations "let x be a in e" == "CONST let1(a, \x. e)" @@ -65,9 +65,9 @@ \g'. f(\x y z. g'(>)))" syntax - "_letrec" :: "[idt,idt,i,i]\i" ("(3letrec _ _ be _/ in _)" [0,0,0,60] 60) - "_letrec2" :: "[idt,idt,idt,i,i]\i" ("(3letrec _ _ _ be _/ in _)" [0,0,0,0,60] 60) - "_letrec3" :: "[idt,idt,idt,idt,i,i]\i" ("(3letrec _ _ _ _ be _/ in _)" [0,0,0,0,0,60] 60) + "_letrec" :: "[idt,idt,i,i]\i" (\(3letrec _ _ be _/ in _)\ [0,0,0,60] 60) + "_letrec2" :: "[idt,idt,idt,i,i]\i" (\(3letrec _ _ _ be _/ in _)\ [0,0,0,0,60] 60) + "_letrec3" :: "[idt,idt,idt,idt,i,i]\i" (\(3letrec _ _ _ _ be _/ in _)\ [0,0,0,0,0,60] 60) syntax_consts "_letrec" == letrec and "_letrec2" == letrec2 and @@ -131,10 +131,10 @@ definition nrec :: "[i,i,[i,i]\i]\i" where "nrec(n,b,c) == letrec g x be ncase(x, b, \y. c(y,g(y))) in g(n)" -definition nil :: "i" ("([])") +definition nil :: "i" (\([])\) where "[] == inl(one)" -definition cons :: "[i,i]\i" (infixr "$" 80) +definition cons :: "[i,i]\i" (infixr \$\ 80) where "h$t == inr()" definition lcase :: "[i,i,[i,i]\i]\i" @@ -143,7 +143,7 @@ definition lrec :: "[i,i,[i,i,i]\i]\i" where "lrec(l,b,c) == letrec g x be lcase(x, b, \h t. c(h,t,g(t))) in g(l)" -definition napply :: "[i\i,i,i]\i" ("(_ ^ _ ` _)" [56,56,56] 56) +definition napply :: "[i\i,i,i]\i" (\(_ ^ _ ` _)\ [56,56,56] 56) where "f ^n` a == nrec(n,a,\x g. f(g))" lemmas simp_can_defs = one_def inl_def inr_def diff -r 46f59511b7bb -r d97fdabd9e2b src/CCL/Trancl.thy --- a/src/CCL/Trancl.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CCL/Trancl.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,13 +15,13 @@ definition id :: "i set" (*the identity relation*) where "id == {p. EX x. p = }" -definition relcomp :: "[i set,i set] \ i set" (infixr "O" 60) (*composition of relations*) +definition relcomp :: "[i set,i set] \ i set" (infixr \O\ 60) (*composition of relations*) where "r O s == {xz. EX x y z. xz = \ :s \ :r}" -definition rtrancl :: "i set \ i set" ("(_^*)" [100] 100) +definition rtrancl :: "i set \ i set" (\(_^*)\ [100] 100) where "r^* == lfp(\s. id Un (r O s))" -definition trancl :: "i set \ i set" ("(_^+)" [100] 100) +definition trancl :: "i set \ i set" (\(_^+)\ [100] 100) where "r^+ == r O rtrancl(r)" diff -r 46f59511b7bb -r d97fdabd9e2b src/CCL/Type.thy --- a/src/CCL/Type.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CCL/Type.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ where "Subtype(A, P) == {x. x:A \ P(x)}" syntax - "_Subtype" :: "[idt, 'a set, o] \ 'a set" ("(1{_: _ ./ _})") + "_Subtype" :: "[idt, 'a set, o] \ 'a set" (\(1{_: _ ./ _})\) syntax_consts "_Subtype" == Subtype translations @@ -25,7 +25,7 @@ definition Bool :: "i set" where "Bool == {x. x=true | x=false}" -definition Plus :: "[i set, i set] \ i set" (infixr "+" 55) +definition Plus :: "[i set, i set] \ i set" (infixr \+\ 55) where "A+B == {x. (EX a:A. x=inl(a)) | (EX b:B. x=inr(b))}" definition Pi :: "[i set, i \ i set] \ i set" @@ -35,10 +35,10 @@ where "Sigma(A,B) == {x. EX a:A. EX b:B(a).x=}" syntax - "_Pi" :: "[idt, i set, i set] \ i set" ("(3PROD _:_./ _)" [0,0,60] 60) - "_Sigma" :: "[idt, i set, i set] \ i set" ("(3SUM _:_./ _)" [0,0,60] 60) - "_arrow" :: "[i set, i set] \ i set" ("(_ ->/ _)" [54, 53] 53) - "_star" :: "[i set, i set] \ i set" ("(_ */ _)" [56, 55] 55) + "_Pi" :: "[idt, i set, i set] \ i set" (\(3PROD _:_./ _)\ [0,0,60] 60) + "_Sigma" :: "[idt, i set, i set] \ i set" (\(3SUM _:_./ _)\ [0,0,60] 60) + "_arrow" :: "[i set, i set] \ i set" (\(_ ->/ _)\ [54, 53] 53) + "_star" :: "[i set, i set] \ i set" (\(_ */ _)\ [56, 55] 55) syntax_consts "_Pi" "_arrow" \ Pi and "_Sigma" "_star" \ Sigma @@ -67,13 +67,13 @@ where "ILists(A) == gfp(\X.{} + A*X)" -definition TAll :: "(i set \ i set) \ i set" (binder "TALL " 55) +definition TAll :: "(i set \ i set) \ i set" (binder \TALL \ 55) where "TALL X. B(X) == Inter({X. EX Y. X=B(Y)})" -definition TEx :: "(i set \ i set) \ i set" (binder "TEX " 55) +definition TEx :: "(i set \ i set) \ i set" (binder \TEX \ 55) where "TEX X. B(X) == Union({X. EX Y. X=B(Y)})" -definition Lift :: "i set \ i set" ("(3[_])") +definition Lift :: "i set \ i set" (\(3[_])\) where "[A] == A Un {bot}" definition SPLIT :: "[i, [i, i] \ i set] \ i set" diff -r 46f59511b7bb -r d97fdabd9e2b src/CCL/Wfd.thy --- a/src/CCL/Wfd.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CCL/Wfd.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,7 +18,7 @@ definition wmap :: "[i\i,i set] \ i set" where "wmap(f,R) == {p. EX x y. p= \ : R}" -definition lex :: "[i set,i set] => i set" (infixl "**" 70) +definition lex :: "[i set,i set] => i set" (infixl \**\ 70) where "ra**rb == {p. EX a a' b b'. p = <,> \ ( : ra | (a=a' \ : rb))}" definition NatPR :: "i set" diff -r 46f59511b7bb -r d97fdabd9e2b src/CCL/ex/List.thy --- a/src/CCL/ex/List.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CCL/ex/List.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,13 +12,13 @@ definition map :: "[i\i,i]\i" where "map(f,l) == lrec(l, [], \x xs g. f(x)$g)" -definition comp :: "[i\i,i\i]\i\i" (infixr "\" 55) +definition comp :: "[i\i,i\i]\i\i" (infixr \\\ 55) where "f \ g == (\x. f(g(x)))" -definition append :: "[i,i]\i" (infixr "@" 55) +definition append :: "[i,i]\i" (infixr \@\ 55) where "l @ m == lrec(l, m, \x xs g. x$g)" -axiomatization member :: "[i,i]\i" (infixr "mem" 55) (* FIXME dangling eq *) +axiomatization member :: "[i,i]\i" (infixr \mem\ 55) (* FIXME dangling eq *) where member_ax: "a mem l == lrec(l, false, \h t g. if eq(a,h) then true else g)" definition filter :: "[i,i]\i" diff -r 46f59511b7bb -r d97fdabd9e2b src/CCL/ex/Nat.thy --- a/src/CCL/ex/Nat.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CCL/ex/Nat.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,28 +12,28 @@ definition not :: "i\i" where "not(b) == if b then false else true" -definition add :: "[i,i]\i" (infixr "#+" 60) +definition add :: "[i,i]\i" (infixr \#+\ 60) where "a #+ b == nrec(a, b, \x g. succ(g))" -definition mult :: "[i,i]\i" (infixr "#*" 60) +definition mult :: "[i,i]\i" (infixr \#*\ 60) where "a #* b == nrec(a, zero, \x g. b #+ g)" -definition sub :: "[i,i]\i" (infixr "#-" 60) +definition sub :: "[i,i]\i" (infixr \#-\ 60) where "a #- b == letrec sub x y be ncase(y, x, \yy. ncase(x, zero, \xx. sub(xx,yy))) in sub(a,b)" -definition le :: "[i,i]\i" (infixr "#<=" 60) +definition le :: "[i,i]\i" (infixr \#<=\ 60) where "a #<= b == letrec le x y be ncase(x, true, \xx. ncase(y, false, \yy. le(xx,yy))) in le(a,b)" -definition lt :: "[i,i]\i" (infixr "#<" 60) +definition lt :: "[i,i]\i" (infixr \#<\ 60) where "a #< b == not(b #<= a)" -definition div :: "[i,i]\i" (infixr "##" 60) +definition div :: "[i,i]\i" (infixr \##\ 60) where "a ## b == letrec div x y be if x #< y then zero else succ(div(x#-y,y)) diff -r 46f59511b7bb -r d97fdabd9e2b src/CTT/CTT.thy --- a/src/CTT/CTT.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/CTT/CTT.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,11 +18,11 @@ consts \ \Judgments\ - Type :: "t \ prop" ("(_ type)" [10] 5) - Eqtype :: "[t,t]\prop" ("(_ =/ _)" [10,10] 5) - Elem :: "[i, t]\prop" ("(_ /: _)" [10,10] 5) - Eqelem :: "[i,i,t]\prop" ("(_ =/ _ :/ _)" [10,10,10] 5) - Reduce :: "[i,i]\prop" ("Reduce[_,_]") + Type :: "t \ prop" (\(_ type)\ [10] 5) + Eqtype :: "[t,t]\prop" (\(_ =/ _)\ [10,10] 5) + Elem :: "[i, t]\prop" (\(_ /: _)\ [10,10] 5) + Eqelem :: "[i,i,t]\prop" (\(_ =/ _ :/ _)\ [10,10,10] 5) + Reduce :: "[i,i]\prop" (\Reduce[_,_]\) \ \Types for truth values\ F :: "t" T :: "t" \ \\F\ is empty, \T\ contains one element\ @@ -30,24 +30,24 @@ tt :: "i" \ \Natural numbers\ N :: "t" - Zero :: "i" ("0") + Zero :: "i" (\0\) succ :: "i\i" rec :: "[i, i, [i,i]\i] \ i" \ \Binary sum\ - Plus :: "[t,t]\t" (infixr "+" 40) + Plus :: "[t,t]\t" (infixr \+\ 40) inl :: "i\i" inr :: "i\i" "when" :: "[i, i\i, i\i]\i" \ \General sum and binary product\ Sum :: "[t, i\t]\t" - pair :: "[i,i]\i" ("(1<_,/_>)") + pair :: "[i,i]\i" (\(1<_,/_>)\) fst :: "i\i" snd :: "i\i" split :: "[i, [i,i]\i] \i" \ \General product and function space\ Prod :: "[t, i\t]\t" - lambda :: "(i \ i) \ i" (binder "\<^bold>\" 10) - app :: "[i,i]\i" (infixl "`" 60) + lambda :: "(i \ i) \ i" (binder \\<^bold>\\ 10) + app :: "[i,i]\i" (infixl \`\ 60) \ \Equality type\ Eq :: "[t,i,i]\t" eq :: "i" @@ -56,8 +56,8 @@ must be introduced after the judgment forms.\ syntax - "_PROD" :: "[idt,t,t]\t" ("(3\_:_./ _)" 10) - "_SUM" :: "[idt,t,t]\t" ("(3\_:_./ _)" 10) + "_PROD" :: "[idt,t,t]\t" (\(3\_:_./ _)\ 10) + "_SUM" :: "[idt,t,t]\t" (\(3\_:_./ _)\ 10) syntax_consts "_PROD" \ Prod and "_SUM" \ Sum @@ -65,10 +65,10 @@ "\x:A. B" \ "CONST Prod(A, \x. B)" "\x:A. B" \ "CONST Sum(A, \x. B)" -abbreviation Arrow :: "[t,t]\t" (infixr "\" 30) +abbreviation Arrow :: "[t,t]\t" (infixr \\\ 30) where "A \ B \ \_:A. B" -abbreviation Times :: "[t,t]\t" (infixr "\" 50) +abbreviation Times :: "[t,t]\t" (infixr \\\ 50) where "A \ B \ \_:A. B" text \ @@ -547,22 +547,22 @@ subsection \Arithmetic operators and their definitions\ -definition add :: "[i,i]\i" (infixr "#+" 65) +definition add :: "[i,i]\i" (infixr \#+\ 65) where "a#+b \ rec(a, b, \u v. succ(v))" -definition diff :: "[i,i]\i" (infixr "-" 65) +definition diff :: "[i,i]\i" (infixr \-\ 65) where "a-b \ rec(b, a, \u v. rec(v, 0, \x y. x))" -definition absdiff :: "[i,i]\i" (infixr "|-|" 65) +definition absdiff :: "[i,i]\i" (infixr \|-|\ 65) where "a|-|b \ (a-b) #+ (b-a)" -definition mult :: "[i,i]\i" (infixr "#*" 70) +definition mult :: "[i,i]\i" (infixr \#*\ 70) where "a#*b \ rec(a, 0, \u v. b #+ v)" -definition mod :: "[i,i]\i" (infixr "mod" 70) +definition mod :: "[i,i]\i" (infixr \mod\ 70) where "a mod b \ rec(a, 0, \u v. rec(succ(v) |-| b, 0, \x y. succ(v)))" -definition div :: "[i,i]\i" (infixr "div" 70) +definition div :: "[i,i]\i" (infixr \div\ 70) where "a div b \ rec(a, 0, \u v. rec(succ(u) mod b, succ(v), \x y. v))" lemmas arith_defs = add_def diff_def absdiff_def mult_def mod_def div_def diff -r 46f59511b7bb -r d97fdabd9e2b src/Cube/Cube.thy --- a/src/Cube/Cube.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Cube/Cube.thy Fri Sep 20 19:51:08 2024 +0200 @@ -22,23 +22,23 @@ Trueprop :: "[context, typing] \ prop" and MT_context :: "context" and Context :: "[typing, context] \ context" and - star :: "term" ("*") and - box :: "term" ("\") and - app :: "[term, term] \ term" (infixl "\" 20) and + star :: "term" (\*\) and + box :: "term" (\\\) and + app :: "[term, term] \ term" (infixl \\\ 20) and Has_type :: "[term, term] \ typing" nonterminal context' and typing' syntax - "_Trueprop" :: "[context', typing'] \ prop" ("(_/ \ _)") - "_Trueprop1" :: "typing' \ prop" ("(_)") - "" :: "id \ context'" ("_") - "" :: "var \ context'" ("_") - "_MT_context" :: "context'" ("") - "_Context" :: "[typing', context'] \ context'" ("_ _") - "_Has_type" :: "[term, term] \ typing'" ("(_:/ _)" [0, 0] 5) - "_Lam" :: "[idt, term, term] \ term" ("(3\<^bold>\_:_./ _)" [0, 0, 0] 10) - "_Pi" :: "[idt, term, term] \ term" ("(3\_:_./ _)" [0, 0] 10) - "_arrow" :: "[term, term] \ term" (infixr "\" 10) + "_Trueprop" :: "[context', typing'] \ prop" (\(_/ \ _)\) + "_Trueprop1" :: "typing' \ prop" (\(_)\) + "" :: "id \ context'" (\_\) + "" :: "var \ context'" (\_\) + "_MT_context" :: "context'" (\\) + "_Context" :: "[typing', context'] \ context'" (\_ _\) + "_Has_type" :: "[term, term] \ typing'" (\(_:/ _)\ [0, 0] 5) + "_Lam" :: "[idt, term, term] \ term" (\(3\<^bold>\_:_./ _)\ [0, 0, 0] 10) + "_Pi" :: "[idt, term, term] \ term" (\(3\_:_./ _)\ [0, 0] 10) + "_arrow" :: "[term, term] \ term" (infixr \\\ 10) syntax_consts "_Trueprop" \ Trueprop and "_MT_context" \ MT_context and diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Classes/Setup.thy --- a/src/Doc/Classes/Setup.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Classes/Setup.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,10 +8,10 @@ declare [[default_code_width = 74]] syntax - "_alpha" :: "type" ("\") - "_alpha_ofsort" :: "sort \ type" ("\' ::_" [0] 1000) - "_beta" :: "type" ("\") - "_beta_ofsort" :: "sort \ type" ("\' ::_" [0] 1000) + "_alpha" :: "type" (\\\) + "_alpha_ofsort" :: "sort \ type" (\\' ::_\ [0] 1000) + "_beta" :: "type" (\\\) + "_beta_ofsort" :: "sort \ type" (\\' ::_\ [0] 1000) parse_ast_translation \ let diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Codegen/Introduction.thy --- a/src/Doc/Codegen/Introduction.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Codegen/Introduction.thy Fri Sep 20 19:51:08 2024 +0200 @@ -123,11 +123,11 @@ \ class %quote semigroup = - fixes mult :: "'a \ 'a \ 'a" (infixl "\" 70) + fixes mult :: "'a \ 'a \ 'a" (infixl \\\ 70) assumes assoc: "(x \ y) \ z = x \ (y \ z)" class %quote monoid = semigroup + - fixes neutral :: 'a ("\") + fixes neutral :: 'a (\\\) assumes neutl: "\ \ x = x" and neutr: "x \ \ = x" diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Codegen/Setup.thy --- a/src/Doc/Codegen/Setup.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Codegen/Setup.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,12 +11,12 @@ ML_file \../more_antiquote.ML\ no_syntax (output) - "_constrain" :: "logic => type => logic" ("_::_" [4, 0] 3) - "_constrain" :: "prop' => type => prop'" ("_::_" [4, 0] 3) + "_constrain" :: "logic => type => logic" (\_::_\ [4, 0] 3) + "_constrain" :: "prop' => type => prop'" (\_::_\ [4, 0] 3) syntax (output) - "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3) - "_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3) + "_constrain" :: "logic => type => logic" (\_ :: _\ [4, 0] 3) + "_constrain" :: "prop' => type => prop'" (\_ :: _\ [4, 0] 3) declare [[default_code_width = 74]] diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Datatypes/Datatypes.thy --- a/src/Doc/Datatypes/Datatypes.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Datatypes/Datatypes.thy Fri Sep 20 19:51:08 2024 +0200 @@ -361,8 +361,8 @@ "[x]" == "x # []" no_notation - Nil ("[]") and - Cons (infixr "#" 65) + Nil (\[]\) and + Cons (infixr \#\ 65) hide_type list hide_const Nil Cons case_list hd tl set map list_all2 rec_list size_list list_all @@ -436,13 +436,13 @@ (*<*) end (*>*) - datatype ('a, 'b) prod (infixr "*" 20) = Pair 'a 'b + datatype ('a, 'b) prod (infixr \*\ 20) = Pair 'a 'b text \\blankline\ datatype (set: 'a) list = - null: Nil ("[]") - | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65) + null: Nil (\[]\) + | Cons (hd: 'a) (tl: "'a list") (infixr \#\ 65) for map: map rel: list_all2 @@ -453,7 +453,7 @@ Incidentally, this is how the traditional syntax can be set up: \ - syntax "_list" :: "list_args \ 'a list" ("[(_)]") + syntax "_list" :: "list_args \ 'a list" (\[(_)]\) text \\blankline\ diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Isar_Ref/First_Order_Logic.thy --- a/src/Doc/Isar_Ref/First_Order_Logic.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Isar_Ref/First_Order_Logic.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ typedecl i typedecl o -judgment Trueprop :: "o \ prop" ("_" 5) +judgment Trueprop :: "o \ prop" (\_\ 5) text \ Note that the object-logic judgment is implicit in the syntax: writing @@ -35,7 +35,7 @@ \B x\ against expressions containing occurrences of \x\. \ -axiomatization equal :: "i \ i \ o" (infix "=" 50) +axiomatization equal :: "i \ i \ o" (infix \=\ 50) where refl [intro]: "x = x" and subst [elim]: "x = y \ B x \ B y" @@ -87,9 +87,9 @@ \ locale group = - fixes prod :: "i \ i \ i" (infix "\" 70) - and inv :: "i \ i" ("(_\)" [1000] 999) - and unit :: i ("1") + fixes prod :: "i \ i \ i" (infix \\\ 70) + and inv :: "i \ i" (\(_\)\ [1000] 999) + and unit :: i (\1\) assumes assoc: "(x \ y) \ z = x \ (y \ z)" and left_unit: "1 \ x = x" and left_inv: "x\ \ x = 1" @@ -167,16 +167,16 @@ Gentzen's system of Natural Deduction \<^cite>\"Gentzen:1935"\. \ -axiomatization imp :: "o \ o \ o" (infixr "\" 25) +axiomatization imp :: "o \ o \ o" (infixr \\\ 25) where impI [intro]: "(A \ B) \ A \ B" and impD [dest]: "(A \ B) \ A \ B" -axiomatization disj :: "o \ o \ o" (infixr "\" 30) +axiomatization disj :: "o \ o \ o" (infixr \\\ 30) where disjI\<^sub>1 [intro]: "A \ A \ B" and disjI\<^sub>2 [intro]: "B \ A \ B" and disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" -axiomatization conj :: "o \ o \ o" (infixr "\" 35) +axiomatization conj :: "o \ o \ o" (infixr \\\ 35) where conjI [intro]: "A \ B \ A \ B" and conjD\<^sub>1: "A \ B \ A" and conjD\<^sub>2: "A \ B \ B" @@ -223,10 +223,10 @@ that is trivially true. \ -axiomatization false :: o ("\") +axiomatization false :: o (\\\) where falseE [elim]: "\ \ A" -definition true :: o ("\") +definition true :: o (\\\) where "\ \ \ \ \" theorem trueI [intro]: \ @@ -237,7 +237,7 @@ Now negation represents an implication towards absurdity: \ -definition not :: "o \ o" ("\ _" [40] 40) +definition not :: "o \ o" (\\ _\ [40] 40) where "\ A \ A \ \" theorem notI [intro]: @@ -320,11 +320,11 @@ notation turns \All (\x. B x)\ into \\x. B x\ etc. \ -axiomatization All :: "(i \ o) \ o" (binder "\" 10) +axiomatization All :: "(i \ o) \ o" (binder \\\ 10) where allI [intro]: "(\x. B x) \ \x. B x" and allD [dest]: "(\x. B x) \ B a" -axiomatization Ex :: "(i \ o) \ o" (binder "\" 10) +axiomatization Ex :: "(i \ o) \ o" (binder \\\ 10) where exI [intro]: "B a \ (\x. B x)" and exE [elim]: "(\x. B x) \ (\x. B x \ C) \ C" diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Isar_Ref/Generic.thy --- a/src/Doc/Isar_Ref/Generic.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Isar_Ref/Generic.thy Fri Sep 20 19:51:08 2024 +0200 @@ -669,7 +669,7 @@ \ experiment - fixes f :: "'a \ 'a \ 'a" (infix "\" 60) + fixes f :: "'a \ 'a \ 'a" (infix \\\ 60) assumes assoc: "(x \ y) \ z = x \ (y \ z)" assumes commute: "x \ y = y \ x" begin diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Isar_Ref/HOL_Specific.thy --- a/src/Doc/Isar_Ref/HOL_Specific.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Isar_Ref/HOL_Specific.thy Fri Sep 20 19:51:08 2024 +0200 @@ -214,7 +214,7 @@ text \The accessible part of a relation is defined as follows:\ inductive acc :: "('a \ 'a \ bool) \ 'a \ bool" - for r :: "'a \ 'a \ bool" (infix "\" 50) + for r :: "'a \ 'a \ bool" (infix \\\ 50) where acc: "(\y. y \ x \ acc r y) \ acc r x" (*<*)end(*>*) diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Isar_Ref/Synopsis.thy --- a/src/Doc/Isar_Ref/Synopsis.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Isar_Ref/Synopsis.thy Fri Sep 20 19:51:08 2024 +0200 @@ -33,7 +33,7 @@ term "?f ?f" txt \Notation:\ - write x ("***") + write x (\***\) end @@ -562,7 +562,7 @@ notepad begin - write Trueprop ("Tr") + write Trueprop (\Tr\) thm conjI thm impI diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Locales/Examples.thy --- a/src/Doc/Locales/Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Locales/Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -50,7 +50,7 @@ \ locale partial_order = - fixes le :: "'a \ 'a \ bool" (infixl "\" 50) + fixes le :: "'a \ 'a \ bool" (infixl \\\ 50) assumes refl [intro, simp]: "x \ x" and anti_sym [intro]: "\ x \ y; y \ x \ \ x = y" and trans [trans]: "\ x \ y; y \ z \ \ x \ z" @@ -141,7 +141,7 @@ \ definition (in partial_order) - less :: "'a \ 'a \ bool" (infixl "\" 50) + less :: "'a \ 'a \ bool" (infixl \\\ 50) where "(x \ y) = (x \ y \ x \ y)" text (in partial_order) \The strict order \less\ with infix @@ -321,9 +321,9 @@ together with an example theorem.\ definition - meet (infixl "\" 70) where "x \ y = (THE inf. is_inf x y inf)" + meet (infixl \\\ 70) where "x \ y = (THE inf. is_inf x y inf)" definition - join (infixl "\" 65) where "x \ y = (THE sup. is_sup x y sup)" + join (infixl \\\ 65) where "x \ y = (THE sup. is_sup x y sup)" lemma %invisible meet_equality [elim?]: "is_inf x y i \ x \ y = i" proof (unfold meet_def) diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Locales/Examples3.thy --- a/src/Doc/Locales/Examples3.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Locales/Examples3.thy Fri Sep 20 19:51:08 2024 +0200 @@ -204,7 +204,7 @@ locale order_preserving = le: partial_order le + le': partial_order le' - for le (infixl "\" 50) and le' (infixl "\" 50) + + for le (infixl \\\ 50) and le' (infixl \\\ 50) + fixes \ assumes hom_le: "x \ y \ \ x \ \ y" @@ -222,7 +222,7 @@ available for the original instance it was declared for. We may introduce infix syntax for \le'.less\ with the following declaration:\ - notation (in order_preserving) le'.less (infixl "\" 50) + notation (in order_preserving) le'.less (infixl \\\ 50) text (in order_preserving) \Now the theorem is displayed nicely as @{thm [source] le'.less_le_trans}: @@ -279,7 +279,7 @@ join.\ locale lattice_hom = - le: lattice + le': lattice le' for le' (infixl "\" 50) + + le: lattice + le': lattice le' for le' (infixl \\\ 50) + fixes \ assumes hom_meet: "\ (x \ y) = le'.meet (\ x) (\ y)" and hom_join: "\ (x \ y) = le'.join (\ x) (\ y)" @@ -295,8 +295,8 @@ context lattice_hom begin - notation le'.meet (infixl "\''" 50) - notation le'.join (infixl "\''" 50) + notation le'.meet (infixl \\''\ 50) + notation le'.join (infixl \\''\ 50) end text \The next example makes radical use of the short hand diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Prog_Prove/LaTeXsugar.thy --- a/src/Doc/Prog_Prove/LaTeXsugar.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Prog_Prove/LaTeXsugar.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,39 +9,39 @@ begin (* DUMMY *) -consts DUMMY :: 'a ("\<^latex>\\\_\") +consts DUMMY :: 'a (\\<^latex>\\_\\) (* THEOREMS *) notation (Rule output) - Pure.imp ("\<^latex>\\\mbox{}\\inferrule{\\mbox{\_\<^latex>\}}\\<^latex>\{\\mbox{\_\<^latex>\}}\") + Pure.imp (\\<^latex>\\mbox{}\inferrule{\mbox{\_\<^latex>\}}\\<^latex>\{\mbox{\_\<^latex>\}}\\) syntax (Rule output) "_bigimpl" :: "asms \ prop \ prop" - ("\<^latex>\\\mbox{}\\inferrule{\_\<^latex>\}\\<^latex>\{\\mbox{\_\<^latex>\}}\") + (\\<^latex>\\mbox{}\inferrule{\_\<^latex>\}\\<^latex>\{\mbox{\_\<^latex>\}}\\) "_asms" :: "prop \ asms \ asms" - ("\<^latex>\\\mbox{\_\<^latex>\}\\\\\/ _") + (\\<^latex>\\mbox{\_\<^latex>\}\\\/ _\) - "_asm" :: "prop \ asms" ("\<^latex>\\\mbox{\_\<^latex>\}\") + "_asm" :: "prop \ asms" (\\<^latex>\\mbox{\_\<^latex>\}\\) notation (Axiom output) - "Trueprop" ("\<^latex>\\\mbox{}\\inferrule{\\mbox{}}{\\mbox{\_\<^latex>\}}\") + "Trueprop" (\\<^latex>\\mbox{}\inferrule{\mbox{}}{\mbox{\_\<^latex>\}}\\) notation (IfThen output) - Pure.imp ("\<^latex>\{\\normalsize{}\If\<^latex>\\\,}\ _/ \<^latex>\{\\normalsize \\,\then\<^latex>\\\,}\/ _.") + Pure.imp (\\<^latex>\{\normalsize{}\If\<^latex>\\,}\ _/ \<^latex>\{\normalsize \,\then\<^latex>\\,}\/ _.\) syntax (IfThen output) "_bigimpl" :: "asms \ prop \ prop" - ("\<^latex>\{\\normalsize{}\If\<^latex>\\\,}\ _ /\<^latex>\{\\normalsize \\,\then\<^latex>\\\,}\/ _.") - "_asms" :: "prop \ asms \ asms" ("\<^latex>\\\mbox{\_\<^latex>\}\ /\<^latex>\{\\normalsize \\,\and\<^latex>\\\,}\/ _") - "_asm" :: "prop \ asms" ("\<^latex>\\\mbox{\_\<^latex>\}\") + (\\<^latex>\{\normalsize{}\If\<^latex>\\,}\ _ /\<^latex>\{\normalsize \,\then\<^latex>\\,}\/ _.\) + "_asms" :: "prop \ asms \ asms" (\\<^latex>\\mbox{\_\<^latex>\}\ /\<^latex>\{\normalsize \,\and\<^latex>\\,}\/ _\) + "_asm" :: "prop \ asms" (\\<^latex>\\mbox{\_\<^latex>\}\\) notation (IfThenNoBox output) - Pure.imp ("\<^latex>\{\\normalsize{}\If\<^latex>\\\,}\ _/ \<^latex>\{\\normalsize \\,\then\<^latex>\\\,}\/ _.") + Pure.imp (\\<^latex>\{\normalsize{}\If\<^latex>\\,}\ _/ \<^latex>\{\normalsize \,\then\<^latex>\\,}\/ _.\) syntax (IfThenNoBox output) "_bigimpl" :: "asms \ prop \ prop" - ("\<^latex>\{\\normalsize{}\If\<^latex>\\\,}\ _ /\<^latex>\{\\normalsize \\,\then\<^latex>\\\,}\/ _.") - "_asms" :: "prop \ asms \ asms" ("_ /\<^latex>\{\\normalsize \\,\and\<^latex>\\\,}\/ _") - "_asm" :: "prop \ asms" ("_") + (\\<^latex>\{\normalsize{}\If\<^latex>\\,}\ _ /\<^latex>\{\normalsize \,\then\<^latex>\\,}\/ _.\) + "_asms" :: "prop \ asms \ asms" (\_ /\<^latex>\{\normalsize \,\and\<^latex>\\,}\/ _\) + "_asm" :: "prop \ asms" (\_\) setup \ Document_Output.antiquotation_pretty_source \<^binding>\const_typ\ diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/CTL/CTL.thy --- a/src/Doc/Tutorial/CTL/CTL.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/CTL/CTL.thy Fri Sep 20 19:51:08 2024 +0200 @@ -35,7 +35,7 @@ a new datatype and a new function.} \ (*<*) -primrec valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) where +primrec valid :: "state \ formula \ bool" (\(_ \ _)\ [80,80] 80) where "s \ Atom a = (a \ L s)" | "s \ Neg f = (~(s \ f))" | "s \ And f g = (s \ f \ s \ g)" | diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/CTL/PDL.thy --- a/src/Doc/Tutorial/CTL/PDL.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/CTL/PDL.thy Fri Sep 20 19:51:08 2024 +0200 @@ -26,7 +26,7 @@ \hbox{\valid s f\}. The definition is by recursion over the syntax: \ -primrec valid :: "state \ formula \ bool" ("(_ \ _)" [80,80] 80) +primrec valid :: "state \ formula \ bool" (\(_ \ _)\ [80,80] 80) where "s \ Atom a = (a \ L s)" | "s \ Neg f = (\(s \ f))" | diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/Documents/Documents.thy --- a/src/Doc/Tutorial/Documents/Documents.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/Documents/Documents.thy Fri Sep 20 19:51:08 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. diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/Inductive/Star.thy --- a/src/Doc/Tutorial/Inductive/Star.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/Inductive/Star.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,7 +15,7 @@ \ inductive_set - rtc :: "('a \ 'a)set \ ('a \ 'a)set" ("_*" [1000] 999) + rtc :: "('a \ 'a)set \ ('a \ 'a)set" (\_*\ [1000] 999) for r :: "('a \ 'a)set" where rtc_refl[iff]: "(x,x) \ r*" diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/Misc/fakenat.thy --- a/src/Doc/Tutorial/Misc/fakenat.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/Misc/fakenat.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,7 +8,7 @@ It behaves approximately as if it were declared like this: \ -datatype nat = zero ("0") | Suc nat +datatype nat = zero (\0\) | Suc nat (*<*) end (*>*) diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/Protocol/Message.thy --- a/src/Doc/Tutorial/Protocol/Message.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/Protocol/Message.thy Fri Sep 20 19:51:08 2024 +0200 @@ -81,9 +81,9 @@ text\Concrete syntax: messages appear as \\A,B,NA\\, etc...\ nonterminal mtuple_args syntax - "" :: "'a \ mtuple_args" ("_") - "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" ("_,/ _") - "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" ("(2\_,/ _\)") + "" :: "'a \ mtuple_args" (\_\) + "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" (\_,/ _\) + "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" (\(2\_,/ _\)\) syntax_consts "_MTuple_args" "_MTuple" \ MPair translations diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/ToyList/ToyList.thy --- a/src/Doc/Tutorial/ToyList/ToyList.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/ToyList/ToyList.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,12 +9,12 @@ the concrete syntax and name space of theory \<^theory>\Main\ as follows. \ -no_notation Nil ("[]") and Cons (infixr "#" 65) and append (infixr "@" 65) +no_notation Nil (\[]\) and Cons (infixr \#\ 65) and append (infixr \@\ 65) hide_type list hide_const rev -datatype 'a list = Nil ("[]") - | Cons 'a "'a list" (infixr "#" 65) +datatype 'a list = Nil (\[]\) + | Cons 'a "'a list" (infixr \#\ 65) text\\noindent The datatype\index{datatype@\isacommand {datatype} (command)} @@ -47,7 +47,7 @@ in this order, because Isabelle insists on definition before use: \ -primrec app :: "'a list \ 'a list \ 'a list" (infixr "@" 65) where +primrec app :: "'a list \ 'a list \ 'a list" (infixr \@\ 65) where "[] @ ys = ys" | "(x # xs) @ ys = x # (xs @ ys)" diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/Types/Axioms.thy --- a/src/Doc/Tutorial/Types/Axioms.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/Types/Axioms.thy Fri Sep 20 19:51:08 2024 +0200 @@ -84,7 +84,7 @@ parameter \neutral\ together with its property:\ class monoidl = semigroup + - fixes neutral :: "'a" ("\") + fixes neutral :: "'a" (\\\) assumes neutl: "\ \ x = x" text \\noindent Again, we prove some instances, by providing @@ -140,7 +140,7 @@ text \\noindent To finish our small algebra example, we add a \group\ class:\ class group = monoidl + - fixes inv :: "'a \ 'a" ("\
_" [81] 80) + fixes inv :: "'a \ 'a" (\\
_\ [81] 80) assumes invl: "\
x \ x = \" text \\noindent We continue with a further example for abstract diff -r 46f59511b7bb -r d97fdabd9e2b src/Doc/Tutorial/Types/Overloading.thy --- a/src/Doc/Tutorial/Types/Overloading.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Doc/Tutorial/Types/Overloading.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ for arbitrary types by means of a type class:\ class plus = - fixes plus :: "'a \ 'a \ 'a" (infixl "\" 70) + fixes plus :: "'a \ 'a \ 'a" (infixl \\\ 70) text \\noindent This introduces a new class @{class [source] plus}, along with a constant @{const [source] plus} with nice infix syntax. diff -r 46f59511b7bb -r d97fdabd9e2b src/FOL/IFOL.thy --- a/src/FOL/IFOL.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/FOL/IFOL.thy Fri Sep 20 19:51:08 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" (\(2\\<^sub>\\<^sub>1 _./ _)\ [0, 10] 10) syntax_consts "_Uniq" \ Uniq translations "\\<^sub>\\<^sub>1x. P" \ "CONST Uniq (\x. P)" diff -r 46f59511b7bb -r d97fdabd9e2b src/FOLP/IFOLP.thy --- a/src/FOLP/IFOLP.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/FOLP/IFOLP.thy Fri Sep 20 19:51:08 2024 +0200 @@ -23,18 +23,18 @@ consts (*** Judgements ***) Proof :: "[o,p]=>prop" - EqProof :: "[p,p,o]=>prop" ("(3_ /= _ :/ _)" [10,10,10] 5) + EqProof :: "[p,p,o]=>prop" (\(3_ /= _ :/ _)\ [10,10,10] 5) (*** Logical Connectives -- Type Formers ***) - eq :: "['a,'a] => o" (infixl "=" 50) + eq :: "['a,'a] => o" (infixl \=\ 50) True :: "o" False :: "o" - conj :: "[o,o] => o" (infixr "&" 35) - disj :: "[o,o] => o" (infixr "|" 30) - imp :: "[o,o] => o" (infixr "-->" 25) + conj :: "[o,o] => o" (infixr \&\ 35) + disj :: "[o,o] => o" (infixr \|\ 30) + imp :: "[o,o] => o" (infixr \-->\ 25) (*Quantifiers*) - All :: "('a => o) => o" (binder "ALL " 10) - Ex :: "('a => o) => o" (binder "EX " 10) + All :: "('a => o) => o" (binder \ALL \ 10) + Ex :: "('a => o) => o" (binder \EX \ 10) (*Rewriting gadgets*) NORM :: "o => o" norm :: "'a => 'a" @@ -44,23 +44,23 @@ contr :: "p=>p" fst :: "p=>p" snd :: "p=>p" - pair :: "[p,p]=>p" ("(1<_,/_>)") + pair :: "[p,p]=>p" (\(1<_,/_>)\) split :: "[p, [p,p]=>p] =>p" inl :: "p=>p" inr :: "p=>p" "when" :: "[p, p=>p, p=>p]=>p" - lambda :: "(p => p) => p" (binder "lam " 55) - App :: "[p,p]=>p" (infixl "`" 60) - alll :: "['a=>p]=>p" (binder "all " 55) - app :: "[p,'a]=>p" (infixl "^" 55) - exists :: "['a,p]=>p" ("(1[_,/_])") + lambda :: "(p => p) => p" (binder \lam \ 55) + App :: "[p,p]=>p" (infixl \`\ 60) + alll :: "['a=>p]=>p" (binder \all \ 55) + app :: "[p,'a]=>p" (infixl \^\ 55) + exists :: "['a,p]=>p" (\(1[_,/_])\) xsplit :: "[p,['a,p]=>p]=>p" ideq :: "'a=>p" idpeel :: "[p,'a=>p]=>p" nrm :: p NRM :: p -syntax "_Proof" :: "[p,o]=>prop" ("(_ /: _)" [51, 10] 5) +syntax "_Proof" :: "[p,o]=>prop" (\(_ /: _)\ [51, 10] 5) parse_translation \ let fun proof_tr [p, P] = Syntax.const \<^const_syntax>\Proof\ $ P $ p @@ -155,14 +155,14 @@ (**** Definitions ****) -definition Not :: "o => o" ("~ _" [40] 40) +definition Not :: "o => o" (\~ _\ [40] 40) where not_def: "~P == P-->False" -definition iff :: "[o,o] => o" (infixr "<->" 25) +definition iff :: "[o,o] => o" (infixr \<->\ 25) where "P<->Q == (P-->Q) & (Q-->P)" (*Unique existence*) -definition Ex1 :: "('a => o) => o" (binder "EX! " 10) +definition Ex1 :: "('a => o) => o" (binder \EX! \ 10) where ex1_def: "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)" (*Rewriting -- special constants to flag normalized terms and formulae*) diff -r 46f59511b7bb -r d97fdabd9e2b src/FOLP/ex/Nat.thy --- a/src/FOLP/ex/Nat.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/FOLP/ex/Nat.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ instance nat :: "term" .. axiomatization - Zero :: nat ("0") and + Zero :: nat (\0\) and Suc :: "nat => nat" and rec :: "[nat, 'a, [nat, 'a] => 'a] => 'a" and @@ -34,7 +34,7 @@ nrecB0: "b: A ==> nrec(0,b,c) = b : A" and nrecBSuc: "c(n,nrec(n,b,c)) : A ==> nrec(Suc(n),b,c) = c(n,nrec(n,b,c)) : A" -definition add :: "[nat, nat] => nat" (infixl "+" 60) +definition add :: "[nat, nat] => nat" (infixl \+\ 60) where "m + n == rec(m, n, %x y. Suc(y))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/AbelCoset.thy --- a/src/HOL/Algebra/AbelCoset.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/AbelCoset.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,34 +13,34 @@ text \Hiding \<+>\ from \<^theory>\HOL.Sum_Type\ until I come up with better syntax here\ -no_notation Sum_Type.Plus (infixr "<+>" 65) +no_notation Sum_Type.Plus (infixr \<+>\ 65) definition - a_r_coset :: "[_, 'a set, 'a] \ 'a set" (infixl "+>\" 60) + a_r_coset :: "[_, 'a set, 'a] \ 'a set" (infixl \+>\\ 60) where "a_r_coset G = r_coset (add_monoid G)" definition - a_l_coset :: "[_, 'a, 'a set] \ 'a set" (infixl "<+\" 60) + a_l_coset :: "[_, 'a, 'a set] \ 'a set" (infixl \<+\\ 60) where "a_l_coset G = l_coset (add_monoid G)" definition - A_RCOSETS :: "[_, 'a set] \ ('a set)set" ("a'_rcosets\ _" [81] 80) + A_RCOSETS :: "[_, 'a set] \ ('a set)set" (\a'_rcosets\ _\ [81] 80) where "A_RCOSETS G H = RCOSETS (add_monoid G) H" definition - set_add :: "[_, 'a set ,'a set] \ 'a set" (infixl "<+>\" 60) + set_add :: "[_, 'a set ,'a set] \ 'a set" (infixl \<+>\\ 60) where "set_add G = set_mult (add_monoid G)" definition - A_SET_INV :: "[_,'a set] \ 'a set" ("a'_set'_inv\ _" [81] 80) + A_SET_INV :: "[_,'a set] \ 'a set" (\a'_set'_inv\ _\ [81] 80) where "A_SET_INV G H = SET_INV (add_monoid G) H" definition - a_r_congruent :: "[('a,'b)ring_scheme, 'a set] \ ('a*'a)set" ("racong\") + a_r_congruent :: "[('a,'b)ring_scheme, 'a set] \ ('a*'a)set" (\racong\\) where "a_r_congruent G = r_congruent (add_monoid G)" definition - A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \ ('a set) monoid" (infixl "A'_Mod" 65) + A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \ ('a set) monoid" (infixl \A'_Mod\ 65) \ \Actually defined for groups rather than monoids\ where "A_FactGroup G H = FactGroup (add_monoid G) H" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Algebraic_Closure.thy --- a/src/HOL/Algebra/Algebraic_Closure.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Algebraic_Closure.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ subsection \Definitions\ -inductive iso_incl :: "'a ring \ 'a ring \ bool" (infixl "\" 65) for A B +inductive iso_incl :: "'a ring \ 'a ring \ bool" (infixl \\\ 65) for A B where iso_inclI [intro]: "id \ ring_hom A B \ iso_incl A B" definition law_restrict :: "('a, 'b) ring_scheme \ 'a ring" @@ -33,7 +33,7 @@ \ index_free \

(P, i) \ \\<^bsub>(P, i)\<^esub> \ carrier L \ (ring.eval L) (\ P) \\<^bsub>(P, i)\<^esub> = \\<^bsub>L\<^esub>) }" -abbreviation (in ring) restrict_extensions :: "((('a list \ nat) multiset) \ 'a) ring set" ("\") +abbreviation (in ring) restrict_extensions :: "((('a list \ nat) multiset) \ 'a) ring set" (\\\) where "\ \ law_restrict ` extensions" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Algebraic_Closure_Type.thy --- a/src/HOL/Algebra/Algebraic_Closure_Type.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Algebraic_Closure_Type.thy Fri Sep 20 19:51:08 2024 +0200 @@ -468,15 +468,15 @@ bundle alg_closure_syntax begin -notation to_ac ("_\" [1000] 999) -notation of_ac ("_\" [1000] 999) +notation to_ac (\_\\ [1000] 999) +notation of_ac (\_\\ [1000] 999) end bundle alg_closure_syntax' begin -notation (output) to_ac ("_") -notation (output) of_ac ("_") +notation (output) to_ac (\_\) +notation (output) of_ac (\_\) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Congruence.thy --- a/src/HOL/Algebra/Congruence.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Congruence.thy Fri Sep 20 19:51:08 2024 +0200 @@ -28,42 +28,42 @@ subsection \Structure with Carrier and Equivalence Relation \eq\\ record 'a eq_object = "'a partial_object" + - eq :: "'a \ 'a \ bool" (infixl ".=\" 50) + eq :: "'a \ 'a \ bool" (infixl \.=\\ 50) definition - elem :: "_ \ 'a \ 'a set \ bool" (infixl ".\\" 50) + elem :: "_ \ 'a \ 'a set \ bool" (infixl \.\\\ 50) where "x .\\<^bsub>S\<^esub> A \ (\y \ A. x .=\<^bsub>S\<^esub> y)" definition - set_eq :: "_ \ 'a set \ 'a set \ bool" (infixl "{.=}\" 50) + set_eq :: "_ \ 'a set \ 'a set \ bool" (infixl \{.=}\\ 50) where "A {.=}\<^bsub>S\<^esub> B \ ((\x \ A. x .\\<^bsub>S\<^esub> B) \ (\x \ B. x .\\<^bsub>S\<^esub> A))" definition - eq_class_of :: "_ \ 'a \ 'a set" ("class'_of\") + eq_class_of :: "_ \ 'a \ 'a set" (\class'_of\\) where "class_of\<^bsub>S\<^esub> x = {y \ carrier S. x .=\<^bsub>S\<^esub> y}" definition - eq_classes :: "_ \ ('a set) set" ("classes\") + eq_classes :: "_ \ ('a set) set" (\classes\\) where "classes\<^bsub>S\<^esub> = {class_of\<^bsub>S\<^esub> x | x. x \ carrier S}" definition - eq_closure_of :: "_ \ 'a set \ 'a set" ("closure'_of\") + eq_closure_of :: "_ \ 'a set \ 'a set" (\closure'_of\\) where "closure_of\<^bsub>S\<^esub> A = {y \ carrier S. y .\\<^bsub>S\<^esub> A}" definition - eq_is_closed :: "_ \ 'a set \ bool" ("is'_closed\") + eq_is_closed :: "_ \ 'a set \ bool" (\is'_closed\\) where "is_closed\<^bsub>S\<^esub> A \ A \ carrier S \ closure_of\<^bsub>S\<^esub> A = A" abbreviation - not_eq :: "_ \ 'a \ 'a \ bool" (infixl ".\\" 50) + not_eq :: "_ \ 'a \ 'a \ bool" (infixl \.\\\ 50) where "x .\\<^bsub>S\<^esub> y \ \(x .=\<^bsub>S\<^esub> y)" abbreviation - not_elem :: "_ \ 'a \ 'a set \ bool" (infixl ".\\" 50) + not_elem :: "_ \ 'a \ 'a set \ bool" (infixl \.\\\ 50) where "x .\\<^bsub>S\<^esub> A \ \(x .\\<^bsub>S\<^esub> A)" abbreviation - set_not_eq :: "_ \ 'a set \ 'a set \ bool" (infixl "{.\}\" 50) + set_not_eq :: "_ \ 'a set \ 'a set \ bool" (infixl \{.\}\\ 50) where "A {.\}\<^bsub>S\<^esub> B \ \(A {.=}\<^bsub>S\<^esub> B)" locale equivalence = diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Coset.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,23 +11,23 @@ section \Cosets and Quotient Groups\ definition - r_coset :: "[_, 'a set, 'a] \ 'a set" (infixl "#>\" 60) + r_coset :: "[_, 'a set, 'a] \ 'a set" (infixl \#>\\ 60) where "H #>\<^bsub>G\<^esub> a = (\h\H. {h \\<^bsub>G\<^esub> a})" definition - l_coset :: "[_, 'a, 'a set] \ 'a set" (infixl "<#\" 60) + l_coset :: "[_, 'a, 'a set] \ 'a set" (infixl \<#\\ 60) where "a <#\<^bsub>G\<^esub> H = (\h\H. {a \\<^bsub>G\<^esub> h})" definition - RCOSETS :: "[_, 'a set] \ ('a set)set" ("rcosets\ _" [81] 80) + RCOSETS :: "[_, 'a set] \ ('a set)set" (\rcosets\ _\ [81] 80) where "rcosets\<^bsub>G\<^esub> H = (\a\carrier G. {H #>\<^bsub>G\<^esub> a})" definition - set_mult :: "[_, 'a set ,'a set] \ 'a set" (infixl "<#>\" 60) + set_mult :: "[_, 'a set ,'a set] \ 'a set" (infixl \<#>\\ 60) where "H <#>\<^bsub>G\<^esub> K = (\h\H. \k\K. {h \\<^bsub>G\<^esub> k})" definition - SET_INV :: "[_,'a set] \ 'a set" ("set'_inv\ _" [81] 80) + SET_INV :: "[_,'a set] \ 'a set" (\set'_inv\ _\ [81] 80) where "set_inv\<^bsub>G\<^esub> H = (\h\H. {inv\<^bsub>G\<^esub> h})" @@ -35,7 +35,7 @@ assumes coset_eq: "(\x \ carrier G. H #> x = x <# H)" abbreviation - normal_rel :: "['a set, ('a, 'b) monoid_scheme] \ bool" (infixl "\" 60) where + normal_rel :: "['a set, ('a, 'b) monoid_scheme] \ bool" (infixl \\\ 60) where "H \ G \ normal H G" lemma (in comm_group) subgroup_imp_normal: "subgroup A G \ A \ G" @@ -659,7 +659,7 @@ subsubsection\An Equivalence Relation\ definition - r_congruent :: "[('a,'b)monoid_scheme, 'a set] \ ('a*'a)set" ("rcong\ _") + r_congruent :: "[('a,'b)monoid_scheme, 'a set] \ ('a*'a)set" (\rcong\ _\) where "rcong\<^bsub>G\<^esub> H = {(x,y). x \ carrier G \ y \ carrier G \ inv\<^bsub>G\<^esub> x \\<^bsub>G\<^esub> y \ H}" @@ -961,7 +961,7 @@ subsection \Quotient Groups: Factorization of a Group\ definition - FactGroup :: "[('a,'b) monoid_scheme, 'a set] \ ('a set) monoid" (infixl "Mod" 65) + FactGroup :: "[('a,'b) monoid_scheme, 'a set] \ ('a set) monoid" (infixl \Mod\ 65) \ \Actually defined for groups rather than monoids\ where "FactGroup G H = \carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Divisibility.thy --- a/src/HOL/Algebra/Divisibility.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Divisibility.thy Fri Sep 20 19:51:08 2024 +0200 @@ -131,10 +131,10 @@ subsubsection \Function definitions\ -definition factor :: "[_, 'a, 'a] \ bool" (infix "divides\" 65) +definition factor :: "[_, 'a, 'a] \ bool" (infix \divides\\ 65) where "a divides\<^bsub>G\<^esub> b \ (\c\carrier G. b = a \\<^bsub>G\<^esub> c)" -definition associated :: "[_, 'a, 'a] \ bool" (infix "\\" 55) +definition associated :: "[_, 'a, 'a] \ bool" (infix \\\\ 55) where "a \\<^bsub>G\<^esub> b \ a divides\<^bsub>G\<^esub> b \ b divides\<^bsub>G\<^esub> a" abbreviation "division_rel G \ \carrier = carrier G, eq = (\\<^bsub>G\<^esub>), le = (divides\<^bsub>G\<^esub>)\" @@ -821,7 +821,7 @@ definition wfactors ::"('a, _) monoid_scheme \ 'a list \ 'a \ bool" where "wfactors G fs a \ (\x \ (set fs). irreducible G x) \ foldr (\\<^bsub>G\<^esub>) fs \\<^bsub>G\<^esub> \\<^bsub>G\<^esub> a" -abbreviation list_assoc :: "('a, _) monoid_scheme \ 'a list \ 'a list \ bool" (infix "[\]\" 44) +abbreviation list_assoc :: "('a, _) monoid_scheme \ 'a list \ 'a list \ bool" (infix \[\]\\ 44) where "list_assoc G \ list_all2 (\\<^bsub>G\<^esub>)" definition essentially_equal :: "('a, _) monoid_scheme \ 'a list \ 'a list \ bool" @@ -1993,11 +1993,11 @@ subsubsection \Definitions\ -definition isgcd :: "[('a,_) monoid_scheme, 'a, 'a, 'a] \ bool" ("(_ gcdof\ _ _)" [81,81,81] 80) +definition isgcd :: "[('a,_) monoid_scheme, 'a, 'a, 'a] \ bool" (\(_ gcdof\ _ _)\ [81,81,81] 80) where "x gcdof\<^bsub>G\<^esub> a b \ x divides\<^bsub>G\<^esub> a \ x divides\<^bsub>G\<^esub> b \ (\y\carrier G. (y divides\<^bsub>G\<^esub> a \ y divides\<^bsub>G\<^esub> b \ y divides\<^bsub>G\<^esub> x))" -definition islcm :: "[_, 'a, 'a, 'a] \ bool" ("(_ lcmof\ _ _)" [81,81,81] 80) +definition islcm :: "[_, 'a, 'a, 'a] \ bool" (\(_ lcmof\ _ _)\ [81,81,81] 80) where "x lcmof\<^bsub>G\<^esub> a b \ a divides\<^bsub>G\<^esub> x \ b divides\<^bsub>G\<^esub> x \ (\y\carrier G. (a divides\<^bsub>G\<^esub> y \ b divides\<^bsub>G\<^esub> y \ x divides\<^bsub>G\<^esub> y))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Embedded_Algebras.thy --- a/src/HOL/Algebra/Embedded_Algebras.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Embedded_Algebras.thy Fri Sep 20 19:51:08 2024 +0200 @@ -38,7 +38,7 @@ abbreviation (in ring) dependent :: "'a set \ 'a list \ bool" where "dependent K U \ \ independent K U" -definition over :: "('a \ 'b) \ 'a \ 'b" (infixl "over" 65) +definition over :: "('a \ 'b) \ 'a \ 'b" (infixl \over\ 65) where "f over a = f a" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Exact_Sequence.thy --- a/src/HOL/Algebra/Exact_Sequence.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Exact_Sequence.thy Fri Sep 20 19:51:08 2024 +0200 @@ -22,7 +22,7 @@ abbreviation exact_seq_arrow :: "('a \ 'a) \ 'a monoid list \ ('a \ 'a) list \ 'a monoid \ 'a monoid list \ ('a \ 'a) list" - ("(3_ / \\ _)" [1000, 60]) + (\(3_ / \\ _)\ [1000, 60]) where "exact_seq_arrow f t G \ (G # (fst t), f # (snd t))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/FiniteProduct.thy --- a/src/HOL/Algebra/FiniteProduct.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/FiniteProduct.thy Fri Sep 20 19:51:08 2024 +0200 @@ -66,7 +66,7 @@ locale LCD = fixes B :: "'b set" and D :: "'a set" - and f :: "'b \ 'a \ 'a" (infixl "\" 70) + and f :: "'b \ 'a \ 'a" (infixl \\\ 70) assumes left_commute: "\x \ B; y \ B; z \ D\ \ x \ (y \ z) = y \ (x \ z)" and f_closed [simp, intro!]: "!!x y. \x \ B; y \ D\ \ f x y \ D" @@ -246,7 +246,7 @@ locale ACeD = fixes D :: "'a set" - and f :: "'a \ 'a \ 'a" (infixl "\" 70) + and f :: "'a \ 'a \ 'a" (infixl \\\ 70) and e :: 'a assumes ident [simp]: "x \ D \ x \ e = x" and commute: "\x \ D; y \ D\ \ x \ y = y \ x" @@ -307,7 +307,7 @@ syntax "_finprod" :: "index \ idt \ 'a set \ 'b \ 'b" - ("(3\__\_. _)" [1000, 0, 51, 10] 10) + (\(3\__\_. _)\ [1000, 0, 51, 10] 10) syntax_consts "_finprod" \ finprod translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Galois_Connection.thy --- a/src/HOL/Algebra/Galois_Connection.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Galois_Connection.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,19 +12,19 @@ subsection \Definition and basic properties\ record ('a, 'b, 'c, 'd) galcon = - orderA :: "('a, 'c) gorder_scheme" ("\\") - orderB :: "('b, 'd) gorder_scheme" ("\\") - lower :: "'a \ 'b" ("\\<^sup>*\") - upper :: "'b \ 'a" ("\\<^sub>*\") + orderA :: "('a, 'c) gorder_scheme" (\\\\) + orderB :: "('b, 'd) gorder_scheme" (\\\\) + lower :: "'a \ 'b" (\\\<^sup>*\\) + upper :: "'b \ 'a" (\\\<^sub>*\\) type_synonym ('a, 'b) galois = "('a, 'b, unit, unit) galcon" abbreviation "inv_galcon G \ \ orderA = inv_gorder \\<^bsub>G\<^esub>, orderB = inv_gorder \\<^bsub>G\<^esub>, lower = upper G, upper = lower G \" -definition comp_galcon :: "('b, 'c) galois \ ('a, 'b) galois \ ('a, 'c) galois" (infixr "\\<^sub>g" 85) +definition comp_galcon :: "('b, 'c) galois \ ('a, 'b) galois \ ('a, 'c) galois" (infixr \\\<^sub>g\ 85) where "G \\<^sub>g F = \ orderA = orderA F, orderB = orderB G, lower = lower G \ lower F, upper = upper F \ upper G \" -definition id_galcon :: "'a gorder \ ('a, 'a) galois" ("I\<^sub>g") where +definition id_galcon :: "'a gorder \ ('a, 'a) galois" (\I\<^sub>g\) where "I\<^sub>g(A) = \ orderA = A, orderB = A, lower = id, upper = id \" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Group.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,11 +18,11 @@ \ record 'a monoid = "'a partial_object" + - mult :: "['a, 'a] \ 'a" (infixl "\\" 70) - one :: 'a ("\\") + mult :: "['a, 'a] \ 'a" (infixl \\\\ 70) + one :: 'a (\\\\) definition - m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" ("inv\ _" [81] 80) + m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" (\inv\ _\ [81] 80) where "inv\<^bsub>G\<^esub> x = (THE y. y \ carrier G \ x \\<^bsub>G\<^esub> y = \\<^bsub>G\<^esub> \ y \\<^bsub>G\<^esub> x = \\<^bsub>G\<^esub>)" definition @@ -353,7 +353,7 @@ subsection \Power\ consts - pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr "[^]\" 75) + pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr \[^]\\ 75) overloading nat_pow == "pow :: [_, 'a, nat] => 'a" begin @@ -793,7 +793,7 @@ subsection \Direct Products\ definition - DirProd :: "_ \ _ \ ('a \ 'b) monoid" (infixr "\\" 80) where + DirProd :: "_ \ _ \ ('a \ 'b) monoid" (infixr \\\\ 80) where "G \\ H = \carrier = carrier G \ carrier H, mult = (\(g, h) (g', h'). (g \\<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h')), @@ -905,7 +905,7 @@ definition iso :: "_ => _ => ('a => 'b) set" where "iso G H = {h. h \ hom G H \ bij_betw h (carrier G) (carrier H)}" -definition is_iso :: "_ \ _ \ bool" (infixr "\" 60) +definition is_iso :: "_ \ _ \ bool" (infixr \\\ 60) where "G \ H = (iso G H \ {})" definition mon where "mon G H = {f \ hom G H. inj_on f (carrier G)}" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Ideal.thy --- a/src/HOL/Algebra/Ideal.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Ideal.thy Fri Sep 20 19:51:08 2024 +0200 @@ -45,7 +45,7 @@ subsubsection (in ring) \Ideals Generated by a Subset of \<^term>\carrier R\\ -definition genideal :: "_ \ 'a set \ 'a set" ("Idl\ _" [80] 79) +definition genideal :: "_ \ 'a set \ 'a set" (\Idl\ _\ [80] 79) where "genideal R S = \{I. ideal I R \ S \ I}" subsubsection \Principal Ideals\ @@ -411,7 +411,7 @@ text \Generation of Principal Ideals in Commutative Rings\ -definition cgenideal :: "_ \ 'a \ 'a set" ("PIdl\ _" [80] 79) +definition cgenideal :: "_ \ 'a \ 'a set" (\PIdl\ _\ [80] 79) where "cgenideal R a = {x \\<^bsub>R\<^esub> a | x. x \ carrier R}" text \genhideal (?) really generates an ideal\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Ideal_Product.thy --- a/src/HOL/Algebra/Ideal_Product.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Ideal_Product.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ text \In this section, we study the structure of the set of ideals of a given ring.\ inductive_set - ideal_prod :: "[ ('a, 'b) ring_scheme, 'a set, 'a set ] \ 'a set" (infixl "\\" 80) + ideal_prod :: "[ ('a, 'b) ring_scheme, 'a set, 'a set ] \ 'a set" (infixl \\\\ 80) for R and I and J (* both I and J are supposed ideals *) where prod: "\ i \ I; j \ J \ \ i \\<^bsub>R\<^esub> j \ ideal_prod R I J" | sum: "\ s1 \ ideal_prod R I J; s2 \ ideal_prod R I J \ \ s1 \\<^bsub>R\<^esub> s2 \ ideal_prod R I J" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Indexed_Polynomials.thy --- a/src/HOL/Algebra/Indexed_Polynomials.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Indexed_Polynomials.thy Fri Sep 20 19:51:08 2024 +0200 @@ -23,13 +23,13 @@ definition (in ring) indexed_const :: "'a \ ('c multiset \ 'a)" where "indexed_const k = (\m. if m = {#} then k else \)" -definition (in ring) indexed_pmult :: "('c multiset \ 'a) \ 'c \ ('c multiset \ 'a)" (infixl "\" 65) +definition (in ring) indexed_pmult :: "('c multiset \ 'a) \ 'c \ ('c multiset \ 'a)" (infixl \\\ 65) where "indexed_pmult P i = (\m. if i \# m then P (m - {# i #}) else \)" -definition (in ring) indexed_padd :: "_ \ _ \ ('c multiset \ 'a)" (infixl "\" 65) +definition (in ring) indexed_padd :: "_ \ _ \ ('c multiset \ 'a)" (infixl \\\ 65) where "indexed_padd P Q = (\m. (P m) \ (Q m))" -definition (in ring) indexed_var :: "'c \ ('c multiset \ 'a)" ("\\") +definition (in ring) indexed_var :: "'c \ ('c multiset \ 'a)" (\\\\) where "indexed_var i = (indexed_const \) \ i" definition (in ring) index_free :: "('c multiset \ 'a) \ 'c \ bool" @@ -38,7 +38,7 @@ definition (in ring) carrier_coeff :: "('c multiset \ 'a) \ bool" where "carrier_coeff P \ (\m. P m \ carrier R)" -inductive_set (in ring) indexed_pset :: "'c set \ 'a set \ ('c multiset \ 'a) set" ("_ [\\]" 80) +inductive_set (in ring) indexed_pset :: "'c set \ 'a set \ ('c multiset \ 'a) set" (\_ [\\]\ 80) for I and K where indexed_const: "k \ K \ indexed_const k \ (K[\\<^bsub>I\<^esub>])" | indexed_padd: "\ P \ (K[\\<^bsub>I\<^esub>]); Q \ (K[\\<^bsub>I\<^esub>]) \ \ P \ Q \ (K[\\<^bsub>I\<^esub>])" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/IntRing.thy --- a/src/HOL/Algebra/IntRing.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/IntRing.thy Fri Sep 20 19:51:08 2024 +0200 @@ -19,7 +19,7 @@ subsection \\\\: The Set of Integers as Algebraic Structure\ -abbreviation int_ring :: "int ring" ("\") +abbreviation int_ring :: "int ring" (\\\) where "int_ring \ \carrier = UNIV, mult = (*), one = 1, zero = 0, add = (+)\" lemma int_Zcarr [intro!, simp]: "k \ carrier \" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Lattice.thy --- a/src/HOL/Algebra/Lattice.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Lattice.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,11 +15,11 @@ subsection \Supremum and infimum\ definition - sup :: "[_, 'a set] => 'a" ("\\_" [90] 90) + sup :: "[_, 'a set] => 'a" (\\\_\ [90] 90) where "\\<^bsub>L\<^esub>A = (SOME x. least L x (Upper L A))" definition - inf :: "[_, 'a set] => 'a" ("\\_" [90] 90) + inf :: "[_, 'a set] => 'a" (\\\_\ [90] 90) where "\\<^bsub>L\<^esub>A = (SOME x. greatest L x (Lower L A))" definition supr :: @@ -31,10 +31,10 @@ where "infi L A f = \\<^bsub>L\<^esub>(f ` A)" syntax - "_inf1" :: "('a, 'b) gorder_scheme \ pttrns \ 'a \ 'a" ("(3IINF\ _./ _)" [0, 10] 10) - "_inf" :: "('a, 'b) gorder_scheme \ pttrn \ 'c set \ 'a \ 'a" ("(3IINF\ _:_./ _)" [0, 0, 10] 10) - "_sup1" :: "('a, 'b) gorder_scheme \ pttrns \ 'a \ 'a" ("(3SSUP\ _./ _)" [0, 10] 10) - "_sup" :: "('a, 'b) gorder_scheme \ pttrn \ 'c set \ 'a \ 'a" ("(3SSUP\ _:_./ _)" [0, 0, 10] 10) + "_inf1" :: "('a, 'b) gorder_scheme \ pttrns \ 'a \ 'a" (\(3IINF\ _./ _)\ [0, 10] 10) + "_inf" :: "('a, 'b) gorder_scheme \ pttrn \ 'c set \ 'a \ 'a" (\(3IINF\ _:_./ _)\ [0, 0, 10] 10) + "_sup1" :: "('a, 'b) gorder_scheme \ pttrns \ 'a \ 'a" (\(3SSUP\ _./ _)\ [0, 10] 10) + "_sup" :: "('a, 'b) gorder_scheme \ pttrn \ 'c set \ 'a \ 'a" (\(3SSUP\ _:_./ _)\ [0, 0, 10] 10) syntax_consts "_inf1" "_inf" == infi and @@ -47,19 +47,19 @@ "SSUP\<^bsub>L\<^esub> x:A. B" == "CONST supr L A (%x. B)" definition - join :: "[_, 'a, 'a] => 'a" (infixl "\\" 65) + join :: "[_, 'a, 'a] => 'a" (infixl \\\\ 65) where "x \\<^bsub>L\<^esub> y = \\<^bsub>L\<^esub>{x, y}" definition - meet :: "[_, 'a, 'a] => 'a" (infixl "\\" 70) + meet :: "[_, 'a, 'a] => 'a" (infixl \\\\ 70) where "x \\<^bsub>L\<^esub> y = \\<^bsub>L\<^esub>{x, y}" definition - LEAST_FP :: "('a, 'b) gorder_scheme \ ('a \ 'a) \ 'a" ("LFP\") where + LEAST_FP :: "('a, 'b) gorder_scheme \ ('a \ 'a) \ 'a" (\LFP\\) where "LEAST_FP L f = \\<^bsub>L\<^esub> {u \ carrier L. f u \\<^bsub>L\<^esub> u}" \ \least fixed point\ definition - GREATEST_FP:: "('a, 'b) gorder_scheme \ ('a \ 'a) \ 'a" ("GFP\") where + GREATEST_FP:: "('a, 'b) gorder_scheme \ ('a \ 'a) \ 'a" (\GFP\\) where "GREATEST_FP L f = \\<^bsub>L\<^esub> {u \ carrier L. u \\<^bsub>L\<^esub> f u}" \ \greatest fixed point\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Left_Coset.thy --- a/src/HOL/Algebra/Left_Coset.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Left_Coset.thy Fri Sep 20 19:51:08 2024 +0200 @@ -7,11 +7,11 @@ begin definition - LCOSETS :: "[_, 'a set] \ ('a set)set" ("lcosets\ _" [81] 80) + LCOSETS :: "[_, 'a set] \ ('a set)set" (\lcosets\ _\ [81] 80) where "lcosets\<^bsub>G\<^esub> H = (\a\carrier G. {a <#\<^bsub>G\<^esub> H})" definition - LFactGroup :: "[('a,'b) monoid_scheme, 'a set] \ ('a set) monoid" (infixl "LMod" 65) + LFactGroup :: "[('a,'b) monoid_scheme, 'a set] \ ('a set) monoid" (infixl \LMod\ 65) \ \Actually defined for groups rather than monoids\ where "LFactGroup G H = \carrier = lcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Module.thy --- a/src/HOL/Algebra/Module.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Module.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ subsection \Definitions\ record ('a, 'b) module = "'b ring" + - smult :: "['a, 'b] => 'b" (infixl "\\" 70) + smult :: "['a, 'b] => 'b" (infixl \\\\ 70) locale module = R?: cring + M?: abelian_group M for M (structure) + assumes smult_closed [simp, intro]: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Multiplicative_Group.thy --- a/src/HOL/Algebra/Multiplicative_Group.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Multiplicative_Group.thy Fri Sep 20 19:51:08 2024 +0200 @@ -144,7 +144,7 @@ where "phi' m = card {x. 1 \ x \ x \ m \ coprime x m}" notation (latex output) - phi' ("\ _") + phi' (\\ _\) lemma phi'_nonzero: assumes "m > 0" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Order.thy --- a/src/HOL/Algebra/Order.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Order.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ subsection \Partial Orders\ record 'a gorder = "'a eq_object" + - le :: "['a, 'a] => bool" (infixl "\\" 50) + le :: "['a, 'a] => bool" (infixl \\\\ 50) abbreviation inv_gorder :: "_ \ 'a gorder" where "inv_gorder L \ @@ -40,7 +40,7 @@ x \ z \ y \ w" definition - lless :: "[_, 'a, 'a] => bool" (infixl "\\" 50) + lless :: "[_, 'a, 'a] => bool" (infixl \\\\ 50) where "x \\<^bsub>L\<^esub> y \ x \\<^bsub>L\<^esub> y \ x .\\<^bsub>L\<^esub> y" subsubsection \The order relation\ @@ -411,7 +411,7 @@ subsubsection \Intervals\ definition - at_least_at_most :: "('a, 'c) gorder_scheme \ 'a => 'a => 'a set" ("(1\_.._\\)") + at_least_at_most :: "('a, 'c) gorder_scheme \ 'a => 'a => 'a set" (\(1\_.._\\)\) where "\l..u\\<^bsub>A\<^esub> = {x \ carrier A. l \\<^bsub>A\<^esub> x \ x \\<^bsub>A\<^esub> u}" context weak_partial_order @@ -452,7 +452,7 @@ shows "isotone L1 L2 f" using assms by (auto simp add:isotone_def) -abbreviation Monotone :: "('a, 'b) gorder_scheme \ ('a \ 'a) \ bool" ("Mono\") +abbreviation Monotone :: "('a, 'b) gorder_scheme \ ('a \ 'a) \ bool" (\Mono\\) where "Monotone L f \ isotone L L f" lemma use_iso1: @@ -478,7 +478,7 @@ subsubsection \Idempotent functions\ definition idempotent :: - "('a, 'b) gorder_scheme \ ('a \ 'a) \ bool" ("Idem\") where + "('a, 'b) gorder_scheme \ ('a \ 'a) \ bool" (\Idem\\) where "idempotent L f \ \x\carrier L. f (f x) .=\<^bsub>L\<^esub> f x" lemma (in weak_partial_order) idempotent: @@ -568,11 +568,11 @@ subsection \Bounded Orders\ definition - top :: "_ => 'a" ("\\") where + top :: "_ => 'a" (\\\\) where "\\<^bsub>L\<^esub> = (SOME x. greatest L x (carrier L))" definition - bottom :: "_ => 'a" ("\\") where + bottom :: "_ => 'a" (\\\\) where "\\<^bsub>L\<^esub> = (SOME x. least L x (carrier L))" locale weak_partial_order_bottom = weak_partial_order L for L (structure) + diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Polynomial_Divisibility.thy --- a/src/HOL/Algebra/Polynomial_Divisibility.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Polynomial_Divisibility.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,16 +14,16 @@ abbreviation poly_ring :: "_ \ ('a list) ring" where "poly_ring R \ univ_poly R (carrier R)" -abbreviation pirreducible :: "_ \ 'a set \ 'a list \ bool" ("pirreducible\") +abbreviation pirreducible :: "_ \ 'a set \ 'a list \ bool" (\pirreducible\\) where "pirreducible\<^bsub>R\<^esub> K p \ ring_irreducible\<^bsub>(univ_poly R K)\<^esub> p" -abbreviation pprime :: "_ \ 'a set \ 'a list \ bool" ("pprime\") +abbreviation pprime :: "_ \ 'a set \ 'a list \ bool" (\pprime\\) where "pprime\<^bsub>R\<^esub> K p \ ring_prime\<^bsub>(univ_poly R K)\<^esub> p" -definition pdivides :: "_ \ 'a list \ 'a list \ bool" (infix "pdivides\" 65) +definition pdivides :: "_ \ 'a list \ 'a list \ bool" (infix \pdivides\\ 65) where "p pdivides\<^bsub>R\<^esub> q = p divides\<^bsub>(univ_poly R (carrier R))\<^esub> q" -definition rupture :: "_ \ 'a set \ 'a list \ (('a list) set) ring" ("Rupt\") +definition rupture :: "_ \ 'a set \ 'a list \ (('a list) set) ring" (\Rupt\\) where "Rupt\<^bsub>R\<^esub> K p = (K[X]\<^bsub>R\<^esub>) Quot (PIdl\<^bsub>K[X]\<^bsub>R\<^esub>\<^esub> p)" abbreviation (in ring) rupture_surj :: "'a set \ 'a list \ 'a list \ ('a list) set" @@ -309,10 +309,10 @@ definition (in ring) long_division :: "'a list \ 'a list \ ('a list \ 'a list)" where "long_division p q = (THE t. long_divides p q t)" -definition (in ring) pdiv :: "'a list \ 'a list \ 'a list" (infixl "pdiv" 65) +definition (in ring) pdiv :: "'a list \ 'a list \ 'a list" (infixl \pdiv\ 65) where "p pdiv q = (if q = [] then [] else fst (long_division p q))" -definition (in ring) pmod :: "'a list \ 'a list \ 'a list" (infixl "pmod" 65) +definition (in ring) pmod :: "'a list \ 'a list \ 'a list" (infixl \pmod\ 65) where "p pmod q = (if q = [] then p else snd (long_division p q))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Polynomials.thy --- a/src/HOL/Algebra/Polynomials.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Polynomials.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,7 +17,7 @@ abbreviation degree :: "'a list \ nat" where "degree p \ length p - 1" -definition polynomial :: "_ \ 'a set \ 'a list \ bool" ("polynomial\") +definition polynomial :: "_ \ 'a set \ 'a list \ bool" (\polynomial\\) where "polynomial\<^bsub>R\<^esub> K p \ p = [] \ (set p \ K \ lead_coeff p \ \\<^bsub>R\<^esub>)" definition (in ring) monom :: "'a \ nat \ 'a list" @@ -1359,7 +1359,7 @@ subsection \Algebraic Structure of Polynomials\ -definition univ_poly :: "('a, 'b) ring_scheme \'a set \ ('a list) ring" ("_ [X]\" 80) +definition univ_poly :: "('a, 'b) ring_scheme \'a set \ ('a list) ring" (\_ [X]\\ 80) where "univ_poly R K = \ carrier = { p. polynomial\<^bsub>R\<^esub> K p }, mult = ring.poly_mult R, @@ -2111,7 +2111,7 @@ subsection \The X Variable\ -definition var :: "_ \ 'a list" ("X\") +definition var :: "_ \ 'a list" (\X\\) where "X\<^bsub>R\<^esub> = [ \\<^bsub>R\<^esub>, \\<^bsub>R\<^esub> ]" lemma (in ring) eval_var: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/QuotRing.thy --- a/src/HOL/Algebra/QuotRing.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/QuotRing.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ subsection \Multiplication on Cosets\ definition rcoset_mult :: "[('a, _) ring_scheme, 'a set, 'a set, 'a set] \ 'a set" - ("[mod _:] _ \\ _" [81,81,81] 80) + (\[mod _:] _ \\ _\ [81,81,81] 80) where "rcoset_mult R I A B = (\a\A. \b\B. I +>\<^bsub>R\<^esub> (a \\<^bsub>R\<^esub> b))" @@ -47,7 +47,7 @@ subsection \Quotient Ring Definition\ definition FactRing :: "[('a,'b) ring_scheme, 'a set] \ ('a set) ring" - (infixl "Quot" 65) + (infixl \Quot\ 65) where "FactRing R I = \carrier = a_rcosets\<^bsub>R\<^esub> I, mult = rcoset_mult R I, one = (I +>\<^bsub>R\<^esub> \\<^bsub>R\<^esub>), zero = I, add = set_add R\" @@ -539,7 +539,7 @@ where "ring_iso R S = { h. h \ ring_hom R S \ bij_betw h (carrier R) (carrier S) }" definition - is_ring_iso :: "_ \ _ \ bool" (infixr "\" 60) + is_ring_iso :: "_ \ _ \ bool" (infixr \\\ 60) where "R \ S = (ring_iso R S \ {})" definition diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Ring.thy --- a/src/HOL/Algebra/Ring.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Ring.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,8 +13,8 @@ subsection \Abelian Groups\ record 'a ring = "'a monoid" + - zero :: 'a ("\\") - add :: "['a, 'a] \ 'a" (infixl "\\" 65) + zero :: 'a (\\\\) + add :: "['a, 'a] \ 'a" (infixl \\\\ 65) abbreviation add_monoid :: "('a, 'm) ring_scheme \ ('a, 'm) monoid_scheme" @@ -23,15 +23,15 @@ text \Derived operations.\ definition - a_inv :: "[('a, 'm) ring_scheme, 'a ] \ 'a" ("\\ _" [81] 80) + a_inv :: "[('a, 'm) ring_scheme, 'a ] \ 'a" (\\\ _\ [81] 80) where "a_inv R = m_inv (add_monoid R)" definition - a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" ("(_ \\ _)" [65,66] 65) + a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" (\(_ \\ _)\ [65,66] 65) where "x \\<^bsub>R\<^esub> y = x \\<^bsub>R\<^esub> (\\<^bsub>R\<^esub> y)" definition - add_pow :: "[_, ('b :: semiring_1), 'a] \ 'a" ("[_] \\ _" [81, 81] 80) + add_pow :: "[_, ('b :: semiring_1), 'a] \ 'a" (\[_] \\ _\ [81, 81] 80) where "add_pow R k a = pow (add_monoid R) a k" locale abelian_monoid = @@ -45,7 +45,7 @@ syntax "_finsum" :: "index \ idt \ 'a set \ 'b \ 'b" - ("(3\__\_. _)" [1000, 0, 51, 10] 10) + (\(3\__\_. _)\ [1000, 0, 51, 10] 10) syntax_consts "_finsum" \ finsum translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Algebra/Ring_Divisibility.thy --- a/src/HOL/Algebra/Ring_Divisibility.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Algebra/Ring_Divisibility.thy Fri Sep 20 19:51:08 2024 +0200 @@ -47,10 +47,10 @@ " \ a \ carrier R - { \ }; b \ carrier R - { \ } \ \ \q r. q \ carrier R \ r \ carrier R \ a = (b \ q) \ r \ ((r = \) \ (\ r < \ b))" -definition ring_irreducible :: "('a, 'b) ring_scheme \ 'a \ bool" ("ring'_irreducible\") +definition ring_irreducible :: "('a, 'b) ring_scheme \ 'a \ bool" (\ring'_irreducible\\) where "ring_irreducible\<^bsub>R\<^esub> a \ (a \ \\<^bsub>R\<^esub>) \ (irreducible R a)" -definition ring_prime :: "('a, 'b) ring_scheme \ 'a \ bool" ("ring'_prime\") +definition ring_prime :: "('a, 'b) ring_scheme \ 'a \ bool" (\ring'_prime\\) where "ring_prime\<^bsub>R\<^esub> a \ (a \ \\<^bsub>R\<^esub>) \ (prime R a)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Abstract_Topology.thy Fri Sep 20 19:51:08 2024 +0200 @@ -607,7 +607,7 @@ subsection\Derived set (set of limit points)\ -definition derived_set_of :: "'a topology \ 'a set \ 'a set" (infixl "derived'_set'_of" 80) +definition derived_set_of :: "'a topology \ 'a set \ 'a set" (infixl \derived'_set'_of\ 80) where "X derived_set_of S \ {x \ topspace X. (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))}" @@ -715,7 +715,7 @@ subsection\ Closure with respect to a topological space\ -definition closure_of :: "'a topology \ 'a set \ 'a set" (infixr "closure'_of" 80) +definition closure_of :: "'a topology \ 'a set \ 'a set" (infixr \closure'_of\ 80) where "X closure_of S \ {x \ topspace X. \T. x \ T \ openin X T \ (\y \ S. y \ T)}" lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \ S)" @@ -947,7 +947,7 @@ text\ Interior with respect to a topological space. \ -definition interior_of :: "'a topology \ 'a set \ 'a set" (infixr "interior'_of" 80) +definition interior_of :: "'a topology \ 'a set \ 'a set" (infixr \interior'_of\ 80) where "X interior_of S \ {x. \T. openin X T \ x \ T \ T \ S}" lemma interior_of_restrict: @@ -1113,7 +1113,7 @@ subsection \Frontier with respect to topological space \ -definition frontier_of :: "'a topology \ 'a set \ 'a set" (infixr "frontier'_of" 80) +definition frontier_of :: "'a topology \ 'a set \ 'a set" (infixr \frontier'_of\ 80) where "X frontier_of S \ X closure_of S - X interior_of S" lemma frontier_of_closures: @@ -3048,7 +3048,7 @@ subsection\Relation of homeomorphism between topological spaces\ -definition homeomorphic_space (infixr "homeomorphic'_space" 50) +definition homeomorphic_space (infixr \homeomorphic'_space\ 50) where "X homeomorphic_space Y \ \f g. homeomorphic_maps X Y f g" lemma homeomorphic_space_refl [iff]: "X homeomorphic_space X" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Abstract_Topology_2.thy --- a/src/HOL/Analysis/Abstract_Topology_2.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Abstract_Topology_2.thy Fri Sep 20 19:51:08 2024 +0200 @@ -311,7 +311,7 @@ subsection\<^marker>\tag unimportant\ \A function constant on a set\ -definition constant_on (infixl "(constant'_on)" 50) +definition constant_on (infixl \(constant'_on)\ 50) where "f constant_on A \ \y. \x\A. f x = y" lemma constant_on_subset: "\f constant_on A; B \ A\ \ f constant_on B" @@ -892,7 +892,7 @@ where "retraction S T r \ T \ S \ continuous_on S r \ r \ S \ T \ (\x\T. r x = x)" -definition\<^marker>\tag important\ retract_of (infixl "retract'_of" 50) where +definition\<^marker>\tag important\ retract_of (infixl \retract'_of\ 50) where "T retract_of S \ (\r. retraction S T r)" lemma retraction_idempotent: "retraction S T r \ x \ S \ r (r x) = r x" @@ -1111,7 +1111,7 @@ subsection\Retractions on a topological space\ -definition retract_of_space :: "'a set \ 'a topology \ bool" (infix "retract'_of'_space" 50) +definition retract_of_space :: "'a set \ 'a topology \ bool" (infix \retract'_of'_space\ 50) where "S retract_of_space X \ S \ topspace X \ (\r. continuous_map X (subtopology X S) r \ (\x \ S. r x = x))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Binary_Product_Measure.thy --- a/src/HOL/Analysis/Binary_Product_Measure.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Binary_Product_Measure.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ subsection "Binary products" -definition\<^marker>\tag important\ pair_measure (infixr "\\<^sub>M" 80) where +definition\<^marker>\tag important\ pair_measure (infixr \\\<^sub>M\ 80) where "A \\<^sub>M B = measure_of (space A \ space B) {a \ b | a b. a \ sets A \ b \ sets B} (\X. \\<^sup>+x. (\\<^sup>+y. indicator X (x,y) \B) \A)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Bochner_Integration.thy Fri Sep 20 19:51:08 2024 +0200 @@ -886,11 +886,11 @@ end -definition\<^marker>\tag important\ lebesgue_integral ("integral\<^sup>L") where +definition\<^marker>\tag important\ lebesgue_integral (\integral\<^sup>L\) where "integral\<^sup>L M f = (if \x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)" syntax - "_lebesgue_integral" :: "pttrn \ real \ 'a measure \ real" ("\((2 _./ _)/ \_)" [60,61] 110) + "_lebesgue_integral" :: "pttrn \ real \ 'a measure \ real" (\\((2 _./ _)/ \_)\ [60,61] 110) syntax_consts "_lebesgue_integral" == lebesgue_integral @@ -899,7 +899,7 @@ "\ x. f \M" == "CONST lebesgue_integral M (\x. f)" syntax - "_ascii_lebesgue_integral" :: "pttrn \ 'a measure \ real \ real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60) + "_ascii_lebesgue_integral" :: "pttrn \ 'a measure \ real \ real" (\(3LINT (1_)/|(_)./ _)\ [0,110,60] 60) syntax_consts "_ascii_lebesgue_integral" == lebesgue_integral diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Borel_Space.thy --- a/src/HOL/Analysis/Borel_Space.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Borel_Space.thy Fri Sep 20 19:51:08 2024 +0200 @@ -859,7 +859,7 @@ by (rule borel_measurable_continuous_Pair) (intro continuous_intros) notation - eucl_less (infix " 50) lemma box_oc: "{x. a x \ b} = {x. a {..b}" and box_co: "{x. a \ x \ x {x. x 50) lemma borel_measurable_Max2[measurable (raw)]: fixes f::"_ \ _ \ 'a::{second_countable_topology, dense_linorder, linorder_topology}" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Bounded_Continuous_Function.thy --- a/src/HOL/Analysis/Bounded_Continuous_Function.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Bounded_Continuous_Function.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,7 +10,7 @@ definition\<^marker>\tag important\ "bcontfun = {f. continuous_on UNIV f \ bounded (range f)}" -typedef (overloaded) ('a, 'b) bcontfun ("(_ \\<^sub>C /_)" [22, 21] 21) = +typedef (overloaded) ('a, 'b) bcontfun (\(_ \\<^sub>C /_)\ [22, 21] 21) = "bcontfun::('a::topological_space \ 'b::metric_space) set" morphisms apply_bcontfun Bcontfun by (auto intro: continuous_intros simp: bounded_def bcontfun_def) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Bounded_Linear_Function.thy --- a/src/HOL/Analysis/Bounded_Linear_Function.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Bounded_Linear_Function.thy Fri Sep 20 19:51:08 2024 +0200 @@ -116,7 +116,7 @@ subsection \Type of bounded linear functions\ -typedef\<^marker>\tag important\ (overloaded) ('a, 'b) blinfun ("(_ \\<^sub>L /_)" [22, 21] 21) = +typedef\<^marker>\tag important\ (overloaded) ('a, 'b) blinfun (\(_ \\<^sub>L /_)\ [22, 21] 21) = "{f::'a::real_normed_vector\'b::real_normed_vector. bounded_linear f}" morphisms blinfun_apply Blinfun by (blast intro: bounded_linear_intros) @@ -663,7 +663,7 @@ lift_definition blinfun_compose:: "'a::real_normed_vector \\<^sub>L 'b::real_normed_vector \ 'c::real_normed_vector \\<^sub>L 'a \ - 'c \\<^sub>L 'b" (infixl "o\<^sub>L" 55) is "(o)" + 'c \\<^sub>L 'b" (infixl \o\<^sub>L\ 55) is "(o)" parametric comp_transfer unfolding o_def by (rule bounded_linear_compose) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Cartesian_Space.thy --- a/src/HOL/Analysis/Cartesian_Space.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Cartesian_Space.thy Fri Sep 20 19:51:08 2024 +0200 @@ -373,8 +373,8 @@ locale linear_first_finite_dimensional_vector_space = l?: Vector_Spaces.linear scaleB scaleC f + B?: finite_dimensional_vector_space scaleB BasisB - for scaleB :: "('a::field => 'b::ab_group_add => 'b)" (infixr "*b" 75) - and scaleC :: "('a => 'c::ab_group_add => 'c)" (infixr "*c" 75) + for scaleB :: "('a::field => 'b::ab_group_add => 'b)" (infixr \*b\ 75) + and scaleC :: "('a => 'c::ab_group_add => 'c)" (infixr \*c\ 75) and BasisB :: "('b set)" and f :: "('b=>'c)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Fri Sep 20 19:51:08 2024 +0200 @@ -159,7 +159,7 @@ subsection\Holomorphic functions\ definition\<^marker>\tag important\ holomorphic_on :: "[complex \ complex, complex set] \ bool" - (infixl "(holomorphic'_on)" 50) + (infixl \(holomorphic'_on)\ 50) where "f holomorphic_on s \ \x\s. f field_differentiable (at x within s)" named_theorems\<^marker>\tag important\ holomorphic_intros "structural introduction rules for holomorphic_on" @@ -375,7 +375,7 @@ subsection\Analyticity on a set\ -definition\<^marker>\tag important\ analytic_on (infixl "(analytic'_on)" 50) +definition\<^marker>\tag important\ analytic_on (infixl \(analytic'_on)\ 50) where "f analytic_on S \ \x \ S. \\. 0 < \ \ f holomorphic_on (ball x \)" named_theorems\<^marker>\tag important\ analytic_intros "introduction rules for proving analyticity" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Cross3.thy --- a/src/HOL/Analysis/Cross3.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Cross3.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ context includes no_Set_Product_syntax begin \\locally disable syntax for set product, to avoid warnings\ -definition\<^marker>\tag important\ cross3 :: "[real^3, real^3] \ real^3" (infixr "\" 80) +definition\<^marker>\tag important\ cross3 :: "[real^3, real^3] \ real^3" (infixr \\\ 80) where "a \ b \ vector [a$2 * b$3 - a$3 * b$2, a$3 * b$1 - a$1 * b$3, @@ -22,13 +22,13 @@ end bundle cross3_syntax begin -notation cross3 (infixr "\" 80) -no_notation Product_Type.Times (infixr "\" 80) +notation cross3 (infixr \\\ 80) +no_notation Product_Type.Times (infixr \\\ 80) end bundle no_cross3_syntax begin -no_notation cross3 (infixr "\" 80) -notation Product_Type.Times (infixr "\" 80) +no_notation cross3 (infixr \\\ 80) +notation Product_Type.Times (infixr \\\ 80) end unbundle cross3_syntax diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Derivative.thy Fri Sep 20 19:51:08 2024 +0200 @@ -100,7 +100,7 @@ definition\<^marker>\tag important\ differentiable_on :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a set \ bool" - (infix "differentiable'_on" 50) + (infix \differentiable'_on\ 50) where "f differentiable_on s \ (\x\s. f differentiable (at x within s))" lemma differentiableI: "(f has_derivative f') net \ f differentiable net" @@ -2008,7 +2008,7 @@ subsection \Field differentiability\ definition\<^marker>\tag important\ field_differentiable :: "['a \ 'a::real_normed_field, 'a filter] \ bool" - (infixr "(field'_differentiable)" 50) + (infixr \(field'_differentiable)\ 50) where "f field_differentiable F \ \f'. (f has_field_derivative f') F" lemma field_differentiable_imp_differentiable: @@ -3089,7 +3089,7 @@ subsection\<^marker>\tag unimportant\ \Piecewise differentiable functions\ definition piecewise_differentiable_on - (infixr "piecewise'_differentiable'_on" 50) + (infixr \piecewise'_differentiable'_on\ 50) where "f piecewise_differentiable_on i \ continuous_on i f \ (\S. finite S \ (\x \ i - S. f differentiable (at x within i)))" @@ -3244,7 +3244,7 @@ asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ definition\<^marker>\tag important\ C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" - (infix "C1'_differentiable'_on" 50) + (infix \C1'_differentiable'_on\ 50) where "f C1_differentiable_on S \ (\D. (\x \ S. (f has_vector_derivative (D x)) (at x)) \ continuous_on S D)" @@ -3336,7 +3336,7 @@ definition\<^marker>\tag important\ piecewise_C1_differentiable_on - (infixr "piecewise'_C1'_differentiable'_on" 50) + (infixr \piecewise'_C1'_differentiable'_on\ 50) where "f piecewise_C1_differentiable_on i \ continuous_on i f \ (\S. finite S \ (f C1_differentiable_on (i - S)))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Elementary_Metric_Spaces.thy --- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Fri Sep 20 19:51:08 2024 +0200 @@ -1553,7 +1553,7 @@ lemma compact_lemma_general: fixes f :: "nat \ 'a" - fixes proj::"'a \ 'b \ 'c::heine_borel" (infixl "proj" 60) + fixes proj::"'a \ 'b \ 'c::heine_borel" (infixl \proj\ 60) fixes unproj:: "('b \ 'c) \ 'a" assumes finite_basis: "finite basis" assumes bounded_proj: "\k. k \ basis \ bounded ((\x. x proj k) ` range f)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Elementary_Normed_Spaces.thy --- a/src/HOL/Analysis/Elementary_Normed_Spaces.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Elementary_Normed_Spaces.thy Fri Sep 20 19:51:08 2024 +0200 @@ -466,7 +466,7 @@ subsection \Filters\ -definition indirection :: "'a::real_normed_vector \ 'a \ 'a filter" (infixr "indirection" 70) +definition indirection :: "'a::real_normed_vector \ 'a \ 'a filter" (infixr \indirection\ 70) where "a indirection v = at a within {b. \c\0. b - a = scaleR c v}" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Elementary_Topology.thy --- a/src/HOL/Analysis/Elementary_Topology.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Elementary_Topology.thy Fri Sep 20 19:51:08 2024 +0200 @@ -563,7 +563,7 @@ subsection \Limit Points\ -definition\<^marker>\tag important\ (in topological_space) islimpt:: "'a \ 'a set \ bool" (infixr "islimpt" 60) +definition\<^marker>\tag important\ (in topological_space) islimpt:: "'a \ 'a set \ bool" (infixr \islimpt\ 60) where "x islimpt S \ (\T. x\T \ open T \ (\y\S. y\T \ y\x))" lemma islimptI: @@ -2221,7 +2221,7 @@ qed definition\<^marker>\tag important\ homeomorphic :: "'a::topological_space set \ 'b::topological_space set \ bool" - (infixr "homeomorphic" 60) + (infixr \homeomorphic\ 60) where "s homeomorphic t \ (\f g. homeomorphism s t f g)" lemma homeomorphic_empty [iff]: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Fri Sep 20 19:51:08 2024 +0200 @@ -964,7 +964,7 @@ abbreviation absolutely_integrable_on :: "('a::euclidean_space \ 'b::{banach, second_countable_topology}) \ 'a set \ bool" - (infixr "absolutely'_integrable'_on" 46) + (infixr \absolutely'_integrable'_on\ 46) where "f absolutely_integrable_on s \ set_integrable lebesgue s f" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy --- a/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Equivalence_Measurable_On_Borel.thy Fri Sep 20 19:51:08 2024 +0200 @@ -480,7 +480,7 @@ subsection\HOL Light measurability\ definition measurable_on :: "('a::euclidean_space \ 'b::real_normed_vector) \ 'a set \ bool" - (infixr "measurable'_on" 46) + (infixr \measurable'_on\ 46) where "f measurable_on S \ \N g. negligible N \ (\n. continuous_on UNIV (g n)) \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Euclidean_Space.thy --- a/src/HOL/Analysis/Euclidean_Space.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Euclidean_Space.thy Fri Sep 20 19:51:08 2024 +0200 @@ -33,7 +33,7 @@ assumes euclidean_all_zero_iff: "(\u\Basis. inner x u = 0) \ (x = 0)" -syntax "_type_dimension" :: "type \ nat" ("(1DIM/(1'(_')))") +syntax "_type_dimension" :: "type \ nat" (\(1DIM/(1'(_')))\) syntax_consts "_type_dimension" \ card translations "DIM('a)" \ "CONST card (CONST Basis :: 'a set)" typed_print_translation \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/FPS_Convergence.thy --- a/src/HOL/Analysis/FPS_Convergence.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/FPS_Convergence.thy Fri Sep 20 19:51:08 2024 +0200 @@ -618,7 +618,7 @@ \ definition\<^marker>\tag important\ has_fps_expansion :: "('a :: {banach,real_normed_div_algebra} \ 'a) \ 'a fps \ bool" - (infixl "has'_fps'_expansion" 60) + (infixl \has'_fps'_expansion\ 60) where "(f has_fps_expansion F) \ fps_conv_radius F > 0 \ eventually (\z. eval_fps F z = f z) (nhds 0)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Analysis/Finite_Cartesian_Product.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Fri Sep 20 19:51:08 2024 +0200 @@ -22,14 +22,14 @@ bundle vec_syntax begin notation - vec_nth (infixl "$" 90) and - vec_lambda (binder "\" 10) + vec_nth (infixl \$\ 90) and + vec_lambda (binder \\\ 10) end bundle no_vec_syntax begin no_notation - vec_nth (infixl "$" 90) and - vec_lambda (binder "\" 10) + vec_nth (infixl \$\ 90) and + vec_lambda (binder \\\ 10) end unbundle vec_syntax @@ -39,7 +39,7 @@ \<^item> \'a^'b\ becomes \('a, 'b::finite) vec\ \<^item> \'a^'b::_\ becomes \('a, 'b) vec\ without extra sort-constraint \ -syntax "_vec_type" :: "type \ type \ type" (infixl "^" 15) +syntax "_vec_type" :: "type \ type \ type" (infixl \^\ 15) syntax_types "_vec_type" \ vec parse_translation \ let @@ -283,7 +283,7 @@ text\Also the scalar-vector multiplication.\ -definition vector_scalar_mult:: "'a::times \ 'a ^ 'n \ 'a ^ 'n" (infixl "*s" 70) +definition vector_scalar_mult:: "'a::times \ 'a ^ 'n \ 'a ^ 'n" (infixl \*s\ 70) where "c *s x = (\ i. c * (x$i))" text \scalar product\ @@ -986,15 +986,15 @@ by (simp add: map_matrix_def) definition\<^marker>\tag important\ matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" - (infixl "**" 70) + (infixl \**\ 70) where "m ** m' == (\ i j. sum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" definition\<^marker>\tag important\ matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" - (infixl "*v" 70) + (infixl \*v\ 70) where "m *v x \ (\ i. sum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" definition\<^marker>\tag important\ vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " - (infixl "v*" 70) + (infixl \v*\ 70) where "v v* m == (\ j. sum (\i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" definition\<^marker>\tag unimportant\ "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Finite_Product_Measure.thy --- a/src/HOL/Analysis/Finite_Product_Measure.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Finite_Product_Measure.thy Fri Sep 20 19:51:08 2024 +0200 @@ -169,7 +169,7 @@ "Pi\<^sub>M I M \ PiM I M" syntax - "_PiM" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" ("(3\\<^sub>M _\_./ _)" 10) + "_PiM" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" (\(3\\<^sub>M _\_./ _)\ 10) syntax_consts "_PiM" == PiM translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Fri Sep 20 19:51:08 2024 +0200 @@ -269,7 +269,7 @@ much later unified. In Fremlin: Measure Theory, Volume 4I this is generalized using residual sets.\ definition has_integral :: "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" - (infixr "has'_integral" 46) + (infixr \has'_integral\ 46) where "(f has_integral I) s \ (if \a b. s = cbox a b then ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter s) @@ -321,7 +321,7 @@ (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" using assms has_integral_alt[of f y i] by auto -definition integrable_on (infixr "integrable'_on" 46) +definition integrable_on (infixr \integrable'_on\ 46) where "f integrable_on i \ (\y. (f has_integral y) i)" definition "integral i f = (SOME y. (f has_integral y) i \ \ f integrable_on i \ y=0)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Homotopy.thy Fri Sep 20 19:51:08 2024 +0200 @@ -3340,7 +3340,7 @@ subsection\Homotopy equivalence of topological spaces.\ definition\<^marker>\tag important\ homotopy_equivalent_space - (infix "homotopy'_equivalent'_space" 50) + (infix \homotopy'_equivalent'_space\ 50) where "X homotopy_equivalent_space Y \ (\f g. continuous_map X Y f \ continuous_map Y X g \ @@ -3781,7 +3781,7 @@ abbreviation\<^marker>\tag important\ homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" - (infix "homotopy'_eqv" 50) + (infix \homotopy'_eqv\ 50) where "S homotopy_eqv T \ top_of_set S homotopy_equivalent_space top_of_set T" lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \ S homotopy_eqv T" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Improper_Integral.thy --- a/src/HOL/Analysis/Improper_Integral.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Improper_Integral.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ text\The definition here only really makes sense for an elementary set. We just use compact intervals in applications below.\ -definition\<^marker>\tag important\ equiintegrable_on (infixr "equiintegrable'_on" 46) +definition\<^marker>\tag important\ equiintegrable_on (infixr \equiintegrable'_on\ 46) where "F equiintegrable_on I \ (\f \ F. f integrable_on I) \ (\e > 0. \\. gauge \ \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Infinite_Products.thy --- a/src/HOL/Analysis/Infinite_Products.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Infinite_Products.thy Fri Sep 20 19:51:08 2024 +0200 @@ -60,14 +60,14 @@ text\The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\ text\<^marker>\tag important\ \%whitespace\ definition\<^marker>\tag important\ - has_prod :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a \ bool" (infixr "has'_prod" 80) + has_prod :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a \ bool" (infixr \has'_prod\ 80) where "f has_prod p \ raw_has_prod f 0 p \ (\i q. p = 0 \ f i = 0 \ raw_has_prod f (Suc i) q)" definition\<^marker>\tag important\ convergent_prod :: "(nat \ 'a :: {t2_space,comm_semiring_1}) \ bool" where "convergent_prod f \ \M p. raw_has_prod f M p" definition\<^marker>\tag important\ prodinf :: "(nat \ 'a::{t2_space, comm_semiring_1}) \ 'a" - (binder "\" 10) + (binder \\\ 10) where "prodinf f = (THE p. f has_prod p)" lemmas prod_defs = raw_has_prod_def has_prod_def convergent_prod_def prodinf_def diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Infinite_Set_Sum.thy --- a/src/HOL/Analysis/Infinite_Set_Sum.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Infinite_Set_Sum.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ *) text \Conflicting notation from \<^theory>\HOL-Analysis.Infinite_Sum\\ -no_notation Infinite_Sum.abs_summable_on (infixr "abs'_summable'_on" 46) +no_notation Infinite_Sum.abs_summable_on (infixr \abs'_summable'_on\ 46) (* TODO Move *) lemma sets_eq_countable: @@ -84,7 +84,7 @@ definition\<^marker>\tag important\ abs_summable_on :: "('a \ 'b :: {banach, second_countable_topology}) \ 'a set \ bool" - (infix "abs'_summable'_on" 50) + (infix \abs'_summable'_on\ 50) where "f abs_summable_on A \ integrable (count_space A) f" @@ -96,10 +96,10 @@ syntax (ASCII) "_infsetsum" :: "pttrn \ 'a set \ 'b \ 'b::{banach, second_countable_topology}" - ("(3INFSETSUM _:_./ _)" [0, 51, 10] 10) + (\(3INFSETSUM _:_./ _)\ [0, 51, 10] 10) syntax "_infsetsum" :: "pttrn \ 'a set \ 'b \ 'b::{banach, second_countable_topology}" - ("(2\\<^sub>a_\_./ _)" [0, 51, 10] 10) + (\(2\\<^sub>a_\_./ _)\ [0, 51, 10] 10) syntax_consts "_infsetsum" \ infsetsum translations \ \Beware of argument permutation!\ @@ -107,10 +107,10 @@ syntax (ASCII) "_uinfsetsum" :: "pttrn \ 'a set \ 'b \ 'b::{banach, second_countable_topology}" - ("(3INFSETSUM _:_./ _)" [0, 51, 10] 10) + (\(3INFSETSUM _:_./ _)\ [0, 51, 10] 10) syntax "_uinfsetsum" :: "pttrn \ 'b \ 'b::{banach, second_countable_topology}" - ("(2\\<^sub>a_./ _)" [0, 10] 10) + (\(2\\<^sub>a_./ _)\ [0, 10] 10) syntax_consts "_uinfsetsum" \ infsetsum translations \ \Beware of argument permutation!\ @@ -118,10 +118,10 @@ syntax (ASCII) "_qinfsetsum" :: "pttrn \ bool \ 'a \ 'a::{banach, second_countable_topology}" - ("(3INFSETSUM _ |/ _./ _)" [0, 0, 10] 10) + (\(3INFSETSUM _ |/ _./ _)\ [0, 0, 10] 10) syntax "_qinfsetsum" :: "pttrn \ bool \ 'a \ 'a::{banach, second_countable_topology}" - ("(2\\<^sub>a_ | (_)./ _)" [0, 0, 10] 10) + (\(2\\<^sub>a_ | (_)./ _)\ [0, 0, 10] 10) syntax_consts "_qinfsetsum" \ infsetsum translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Infinite_Sum.thy --- a/src/HOL/Analysis/Infinite_Sum.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Infinite_Sum.thy Fri Sep 20 19:51:08 2024 +0200 @@ -34,40 +34,40 @@ definition HAS_SUM :: \('a \ 'b :: {comm_monoid_add, topological_space}) \ 'a set \ 'b \ bool\ where has_sum_def: \HAS_SUM f A x \ (sum f \ x) (finite_subsets_at_top A)\ -abbreviation has_sum (infixr "has'_sum" 46) where +abbreviation has_sum (infixr \has'_sum\ 46) where "(f has_sum S) A \ HAS_SUM f A S" -definition summable_on :: "('a \ 'b::{comm_monoid_add, topological_space}) \ 'a set \ bool" (infixr "summable'_on" 46) where +definition summable_on :: "('a \ 'b::{comm_monoid_add, topological_space}) \ 'a set \ bool" (infixr \summable'_on\ 46) where "f summable_on A \ (\x. (f has_sum x) A)" definition infsum :: "('a \ 'b::{comm_monoid_add,t2_space}) \ 'a set \ 'b" where "infsum f A = (if f summable_on A then Lim (finite_subsets_at_top A) (sum f) else 0)" -abbreviation abs_summable_on :: "('a \ 'b::real_normed_vector) \ 'a set \ bool" (infixr "abs'_summable'_on" 46) where +abbreviation abs_summable_on :: "('a \ 'b::real_normed_vector) \ 'a set \ bool" (infixr \abs'_summable'_on\ 46) where "f abs_summable_on A \ (\x. norm (f x)) summable_on A" syntax (ASCII) - "_infsum" :: "pttrn \ 'a set \ 'b \ 'b::topological_comm_monoid_add" ("(3INFSUM (_/:_)./ _)" [0, 51, 10] 10) + "_infsum" :: "pttrn \ 'a set \ 'b \ 'b::topological_comm_monoid_add" (\(3INFSUM (_/:_)./ _)\ [0, 51, 10] 10) syntax - "_infsum" :: "pttrn \ 'a set \ 'b \ 'b::topological_comm_monoid_add" ("(2\\<^sub>\(_/\_)./ _)" [0, 51, 10] 10) + "_infsum" :: "pttrn \ 'a set \ 'b \ 'b::topological_comm_monoid_add" (\(2\\<^sub>\(_/\_)./ _)\ [0, 51, 10] 10) syntax_consts "_infsum" \ infsum translations \ \Beware of argument permutation!\ "\\<^sub>\i\A. b" \ "CONST infsum (\i. b) A" syntax (ASCII) - "_univinfsum" :: "pttrn \ 'a \ 'a" ("(3INFSUM _./ _)" [0, 10] 10) + "_univinfsum" :: "pttrn \ 'a \ 'a" (\(3INFSUM _./ _)\ [0, 10] 10) syntax - "_univinfsum" :: "pttrn \ 'a \ 'a" ("(2\\<^sub>\_./ _)" [0, 10] 10) + "_univinfsum" :: "pttrn \ 'a \ 'a" (\(2\\<^sub>\_./ _)\ [0, 10] 10) syntax_consts "_univinfsum" \ infsum translations "\\<^sub>\x. t" \ "CONST infsum (\x. t) (CONST UNIV)" syntax (ASCII) - "_qinfsum" :: "pttrn \ bool \ 'a \ 'a" ("(3INFSUM _ |/ _./ _)" [0, 0, 10] 10) + "_qinfsum" :: "pttrn \ bool \ 'a \ 'a" (\(3INFSUM _ |/ _./ _)\ [0, 0, 10] 10) syntax - "_qinfsum" :: "pttrn \ bool \ 'a \ 'a" ("(2\\<^sub>\_ | (_)./ _)" [0, 0, 10] 10) + "_qinfsum" :: "pttrn \ bool \ 'a \ 'a" (\(2\\<^sub>\_ | (_)./ _)\ [0, 0, 10] 10) syntax_consts "_qinfsum" \ infsum translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Inner_Product.thy --- a/src/HOL/Analysis/Inner_Product.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Inner_Product.thy Fri Sep 20 19:51:08 2024 +0200 @@ -393,7 +393,7 @@ definition\<^marker>\tag important\ gderiv :: "['a::real_inner \ real, 'a, 'a] \ bool" - ("(GDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) + (\(GDERIV (_)/ (_)/ :> (_))\ [1000, 1000, 60] 60) where "GDERIV f x :> D \ FDERIV f x :> (\h. inner h D)" @@ -463,11 +463,11 @@ lemmas has_derivative_norm = GDERIV_norm [unfolded gderiv_def] bundle inner_syntax begin -notation inner (infix "\" 70) +notation inner (infix \\\ 70) end bundle no_inner_syntax begin -no_notation inner (infix "\" 70) +no_notation inner (infix \\\ 70) end end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Interval_Integral.thy --- a/src/HOL/Analysis/Interval_Integral.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Interval_Integral.thy Fri Sep 20 19:51:08 2024 +0200 @@ -146,7 +146,7 @@ syntax "_ascii_interval_lebesgue_integral" :: "pttrn \ real \ real \ real measure \ real \ real" - ("(5LINT _=_.._|_. _)" [0,60,60,61,100] 60) + (\(5LINT _=_.._|_. _)\ [0,60,60,61,100] 60) syntax_consts "_ascii_interval_lebesgue_integral" == interval_lebesgue_integral @@ -160,7 +160,7 @@ syntax "_ascii_interval_lebesgue_borel_integral" :: "pttrn \ real \ real \ real \ real" - ("(4LBINT _=_.._. _)" [0,60,60,61] 60) + (\(4LBINT _=_.._. _)\ [0,60,60,61] 60) syntax_consts "_ascii_interval_lebesgue_borel_integral" == interval_lebesgue_integral @@ -1049,7 +1049,7 @@ syntax "_complex_lebesgue_borel_integral" :: "pttrn \ real \ complex" - ("(2CLBINT _. _)" [0,60] 60) + (\(2CLBINT _. _)\ [0,60] 60) syntax_consts "_complex_lebesgue_borel_integral" == complex_lebesgue_integral @@ -1057,7 +1057,7 @@ translations "CLBINT x. f" == "CONST complex_lebesgue_integral CONST lborel (\x. f)" syntax "_complex_set_lebesgue_borel_integral" :: "pttrn \ real set \ real \ complex" - ("(3CLBINT _:_. _)" [0,60,61] 60) + (\(3CLBINT _:_. _)\ [0,60,61] 60) syntax_consts "_complex_set_lebesgue_borel_integral" == complex_set_lebesgue_integral @@ -1075,7 +1075,7 @@ syntax "_ascii_complex_interval_lebesgue_borel_integral" :: "pttrn \ ereal \ ereal \ real \ complex" - ("(4CLBINT _=_.._. _)" [0,60,60,61] 60) + (\(4CLBINT _=_.._. _)\ [0,60,60,61] 60) syntax_consts "_ascii_complex_interval_lebesgue_borel_integral" == complex_interval_lebesgue_integral diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Isolated.thy --- a/src/HOL/Analysis/Isolated.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Isolated.thy Fri Sep 20 19:51:08 2024 +0200 @@ -5,7 +5,7 @@ subsection \Isolate and discrete\ -definition (in topological_space) isolated_in:: "'a \ 'a set \ bool" (infixr "isolated'_in" 60) +definition (in topological_space) isolated_in:: "'a \ 'a set \ bool" (infixr \isolated'_in\ 60) where "x isolated_in S \ (x\S \ (\T. open T \ T \ S = {x}))" definition (in topological_space) discrete:: "'a set \ bool" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Lipschitz.thy --- a/src/HOL/Analysis/Lipschitz.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Lipschitz.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,10 +13,10 @@ where "lipschitz_on C U f \ (0 \ C \ (\x \ U. \y\U. dist (f x) (f y) \ C * dist x y))" bundle lipschitz_syntax begin -notation\<^marker>\tag important\ lipschitz_on ("_-lipschitz'_on" [1000]) +notation\<^marker>\tag important\ lipschitz_on (\_-lipschitz'_on\ [1000]) end bundle no_lipschitz_syntax begin -no_notation lipschitz_on ("_-lipschitz'_on" [1000]) +no_notation lipschitz_on (\_-lipschitz'_on\ [1000]) end unbundle lipschitz_syntax diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Locally.thy --- a/src/HOL/Analysis/Locally.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Locally.thy Fri Sep 20 19:51:08 2024 +0200 @@ -945,7 +945,7 @@ text\Basic definition of the small inductive dimension relation. Works in any topological space.\ -inductive dimension_le :: "['a topology, int] \ bool" (infix "dim'_le" 50) +inductive dimension_le :: "['a topology, int] \ bool" (infix \dim'_le\ 50) where "\-1 \ n; \V a. \openin X V; a \ V\ \ \U. a \ U \ U \ V \ openin X U \ (subtopology X (X frontier_of U)) dim_le (n-1)\ \ X dim_le (n::int)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Measure_Space.thy Fri Sep 20 19:51:08 2024 +0200 @@ -1004,7 +1004,7 @@ "almost_everywhere M P \ eventually P (ae_filter M)" syntax - "_almost_everywhere" :: "pttrn \ 'a \ bool \ bool" ("AE _ in _. _" [0,0,10] 10) + "_almost_everywhere" :: "pttrn \ 'a \ bool \ bool" (\AE _ in _. _\ [0,0,10] 10) syntax_consts "_almost_everywhere" \ almost_everywhere @@ -1017,7 +1017,7 @@ syntax "_set_almost_everywhere" :: "pttrn \ 'a set \ 'a \ bool \ bool" - ("AE _\_ in _./ _" [0,0,0,10] 10) + (\AE _\_ in _./ _\ [0,0,0,10] 10) syntax_consts "_set_almost_everywhere" \ set_almost_everywhere diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Fri Sep 20 19:51:08 2024 +0200 @@ -580,11 +580,11 @@ subsection "Simple integral" -definition\<^marker>\tag important\ simple_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>S") where +definition\<^marker>\tag important\ simple_integral :: "'a measure \ ('a \ ennreal) \ ennreal" (\integral\<^sup>S\) where "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" syntax - "_simple_integral" :: "pttrn \ ennreal \ 'a measure \ ennreal" ("\\<^sup>S _. _ \_" [60,61] 110) + "_simple_integral" :: "pttrn \ ennreal \ 'a measure \ ennreal" (\\\<^sup>S _. _ \_\ [60,61] 110) syntax_consts "_simple_integral" == simple_integral @@ -816,11 +816,11 @@ subsection \Integral on nonnegative functions\ -definition\<^marker>\tag important\ nn_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>N") where +definition\<^marker>\tag important\ nn_integral :: "'a measure \ ('a \ ennreal) \ ennreal" (\integral\<^sup>N\) where "integral\<^sup>N M f = (SUP g \ {g. simple_function M g \ g \ f}. integral\<^sup>S M g)" syntax - "_nn_integral" :: "pttrn \ ennreal \ 'a measure \ ennreal" ("\\<^sup>+((2 _./ _)/ \_)" [60,61] 110) + "_nn_integral" :: "pttrn \ ennreal \ 'a measure \ ennreal" (\\\<^sup>+((2 _./ _)/ \_)\ [60,61] 110) syntax_consts "_nn_integral" == nn_integral diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Ordered_Euclidean_Space.thy --- a/src/HOL/Analysis/Ordered_Euclidean_Space.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy Fri Sep 20 19:51:08 2024 +0200 @@ -289,7 +289,7 @@ by (metis compact_cbox interval_cbox) no_notation - eucl_less (infix " 50) lemma One_nonneg: "0 \ (\Basis::'a::ordered_euclidean_space)" by (auto intro: sum_nonneg) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Path_Connected.thy Fri Sep 20 19:51:08 2024 +0200 @@ -28,7 +28,7 @@ where "reversepath g \ (\x. g(1 - x))" definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" - (infixr "+++" 75) + (infixr \+++\ 75) where "g1 +++ g2 \ (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" definition\<^marker>\tag important\ loop_free :: "(real \ 'a::topological_space) \ bool" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Polytope.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,7 +8,7 @@ subsection \Faces of a (usually convex) set\ -definition\<^marker>\tag important\ face_of :: "['a::real_vector set, 'a set] \ bool" (infixr "(face'_of)" 50) +definition\<^marker>\tag important\ face_of :: "['a::real_vector set, 'a set] \ bool" (infixr \(face'_of)\ 50) where "T face_of S \ T \ S \ convex T \ @@ -654,7 +654,7 @@ text\That is, faces that are intersection with supporting hyperplane\ definition\<^marker>\tag important\ exposed_face_of :: "['a::euclidean_space set, 'a set] \ bool" - (infixr "(exposed'_face'_of)" 50) + (infixr \(exposed'_face'_of)\ 50) where "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b})" @@ -873,7 +873,7 @@ subsection\Extreme points of a set: its singleton faces\ definition\<^marker>\tag important\ extreme_point_of :: "['a::real_vector, 'a set] \ bool" - (infixr "(extreme'_point'_of)" 50) + (infixr \(extreme'_point'_of)\ 50) where "x extreme_point_of S \ x \ S \ (\a \ S. \b \ S. x \ open_segment a b)" @@ -979,7 +979,7 @@ subsection\Facets\ definition\<^marker>\tag important\ facet_of :: "['a::euclidean_space set, 'a set] \ bool" - (infixr "(facet'_of)" 50) + (infixr \(facet'_of)\ 50) where "F facet_of S \ F face_of S \ F \ {} \ aff_dim F = aff_dim S - 1" lemma facet_of_empty [simp]: "\ S facet_of {}" @@ -1024,7 +1024,7 @@ subsection \Edges: faces of affine dimension 1\ (*FIXME too small subsection, rearrange? *) -definition\<^marker>\tag important\ edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(edge'_of)" 50) +definition\<^marker>\tag important\ edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr \(edge'_of)\ 50) where "e edge_of S \ e face_of S \ aff_dim e = 1" lemma edge_of_imp_subset: @@ -2984,7 +2984,7 @@ text\The notion of n-simplex for integer \<^term>\n \ -1\\ -definition\<^marker>\tag important\ simplex :: "int \ 'a::euclidean_space set \ bool" (infix "simplex" 50) +definition\<^marker>\tag important\ simplex :: "int \ 'a::euclidean_space set \ bool" (infix \simplex\ 50) where "n simplex S \ \C. \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C" lemma simplex: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Set_Integral.thy --- a/src/HOL/Analysis/Set_Integral.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Set_Integral.thy Fri Sep 20 19:51:08 2024 +0200 @@ -24,7 +24,7 @@ syntax "_ascii_set_lebesgue_integral" :: "pttrn \ 'a set \ 'a measure \ real \ real" - ("(4LINT (_):(_)/|(_)./ _)" [0,0,0,10] 10) + (\(4LINT (_):(_)/|(_)./ _)\ [0,0,0,10] 10) syntax_consts "_ascii_set_lebesgue_integral" == set_lebesgue_integral @@ -43,11 +43,11 @@ syntax "_lebesgue_borel_integral" :: "pttrn \ real \ real" - ("(2LBINT _./ _)" [0,10] 10) + (\(2LBINT _./ _)\ [0,10] 10) syntax "_set_lebesgue_borel_integral" :: "pttrn \ real set \ real \ real" - ("(3LBINT _:_./ _)" [0,0,10] 10) + (\(3LBINT _:_./ _)\ [0,0,10] 10) section \Basic properties\ @@ -572,12 +572,12 @@ abbreviation complex_integrable :: "'a measure \ ('a \ complex) \ bool" where "complex_integrable M f \ integrable M f" -abbreviation complex_lebesgue_integral :: "'a measure \ ('a \ complex) \ complex" ("integral\<^sup>C") where +abbreviation complex_lebesgue_integral :: "'a measure \ ('a \ complex) \ complex" (\integral\<^sup>C\) where "integral\<^sup>C M f == integral\<^sup>L M f" syntax "_complex_lebesgue_integral" :: "pttrn \ complex \ 'a measure \ complex" - ("\\<^sup>C _. _ \_" [0,0] 110) + (\\\<^sup>C _. _ \_\ [0,0] 110) syntax_consts "_complex_lebesgue_integral" == complex_lebesgue_integral @@ -587,7 +587,7 @@ syntax "_ascii_complex_lebesgue_integral" :: "pttrn \ 'a measure \ real \ real" - ("(3CLINT _|_. _)" [0,0,10] 10) + (\(3CLINT _|_. _)\ [0,0,10] 10) syntax_consts "_ascii_complex_lebesgue_integral" == complex_lebesgue_integral @@ -624,7 +624,7 @@ syntax "_ascii_complex_set_lebesgue_integral" :: "pttrn \ 'a set \ 'a measure \ real \ real" - ("(4CLINT _:_|_. _)" [0,0,0,10] 10) + (\(4CLINT _:_|_. _)\ [0,0,0,10] 10) syntax_consts "_ascii_complex_set_lebesgue_integral" == complex_set_lebesgue_integral @@ -647,10 +647,10 @@ syntax "_set_nn_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" -("(\\<^sup>+((_)\(_)./ _)/\_)" [0,0,0,110] 10) +(\(\\<^sup>+((_)\(_)./ _)/\_)\ [0,0,0,110] 10) "_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" -("(\((_)\(_)./ _)/\_)" [0,0,0,110] 10) +(\(\((_)\(_)./ _)/\_)\ [0,0,0,110] 10) syntax_consts "_set_nn_integral" == set_nn_integral and diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Sigma_Algebra.thy --- a/src/HOL/Analysis/Sigma_Algebra.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Sigma_Algebra.thy Fri Sep 20 19:51:08 2024 +0200 @@ -1759,7 +1759,7 @@ subsubsection \Measurable functions\ definition\<^marker>\tag important\ measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" - (infixr "\\<^sub>M" 60) where + (infixr \\\<^sub>M\ 60) where "measurable A B = {f \ space A \ space B. \y \ sets B. f -` y \ space A \ sets A}" lemma measurableI: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Sparse_In.thy --- a/src/HOL/Analysis/Sparse_In.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Sparse_In.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,7 +10,7 @@ subsection \A set of points sparse in another set\ definition sparse_in:: "'a :: topological_space set \ 'a set \ bool" - (infixl "(sparse'_in)" 50) + (infixl \(sparse'_in)\ 50) where "pts sparse_in A = (\x\A. \B. x\B \ open B \ (\y\B. \ y islimpt pts))" @@ -165,14 +165,14 @@ "cosparse A = Abs_filter (\P. {x. \P x} sparse_in A)" syntax - "_eventually_cosparse" :: "pttrn => 'a set => bool => bool" ("(3\\<^sub>\_\_./ _)" [0, 0, 10] 10) + "_eventually_cosparse" :: "pttrn => 'a set => bool => bool" (\(3\\<^sub>\_\_./ _)\ [0, 0, 10] 10) syntax_consts "_eventually_cosparse" == eventually translations "\\<^sub>\x\A. P" == "CONST eventually (\x. P) (CONST cosparse A)" syntax - "_qeventually_cosparse" :: "pttrn \ bool \ 'a \ 'a" ("(3\\<^sub>\_ | (_)./ _)" [0, 0, 10] 10) + "_qeventually_cosparse" :: "pttrn \ bool \ 'a \ 'a" (\(3\\<^sub>\_ | (_)./ _)\ [0, 0, 10] 10) syntax_consts "_qeventually_cosparse" == eventually translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Starlike.thy Fri Sep 20 19:51:08 2024 +0200 @@ -6414,7 +6414,7 @@ subsection\Orthogonal complement\ -definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) +definition\<^marker>\tag important\ orthogonal_comp (\_\<^sup>\\ [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Tagged_Division.thy --- a/src/HOL/Analysis/Tagged_Division.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Tagged_Division.thy Fri Sep 20 19:51:08 2024 +0200 @@ -203,7 +203,7 @@ subsection \Divisions\ -definition\<^marker>\tag important\ division_of (infixl "division'_of" 40) +definition\<^marker>\tag important\ division_of (infixl \division'_of\ 40) where "s division_of i \ finite s \ @@ -820,7 +820,7 @@ subsection \Tagged (partial) divisions\ -definition\<^marker>\tag important\ tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40) +definition\<^marker>\tag important\ tagged_partial_division_of (infixr \tagged'_partial'_division'_of\ 40) where "s tagged_partial_division_of i \ finite s \ (\x K. (x, K) \ s \ x \ K \ K \ i \ (\a b. K = cbox a b)) \ @@ -837,7 +837,7 @@ (x2, K2) \ s \ (x1, K1) \ (x2, K2) \ interior K1 \ interior K2 = {}" using assms unfolding tagged_partial_division_of_def by blast+ -definition\<^marker>\tag important\ tagged_division_of (infixr "tagged'_division'_of" 40) +definition\<^marker>\tag important\ tagged_division_of (infixr \tagged'_division'_of\ 40) where "s tagged_division_of i \ s tagged_partial_division_of i \ (\{K. \x. (x,K) \ s} = i)" lemma tagged_division_of_finite: "s tagged_division_of i \ finite s" @@ -1593,7 +1593,7 @@ subsection \Fine-ness of a partition w.r.t. a gauge\ -definition\<^marker>\tag important\ fine (infixr "fine" 46) +definition\<^marker>\tag important\ fine (infixr \fine\ 46) where "d fine s \ (\(x,k) \ s. k \ d x)" lemma fineI: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Fri Sep 20 19:51:08 2024 +0200 @@ -320,7 +320,7 @@ corollary\<^marker>\tag unimportant\ Zero_neq_One[iff]: "0 \ One" by (metis One_non_0) -definition\<^marker>\tag important\ (in euclidean_space) eucl_less (infix "\tag important\ (in euclidean_space) eucl_less (infix \ 50) where "eucl_less a b \ (\i\Basis. a \ i < b \ i)" definition\<^marker>\tag important\ box_eucl_less: "box a b = {x. a x 50) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/KerberosIV.thy --- a/src/HOL/Auth/KerberosIV.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/KerberosIV.thy Fri Sep 20 19:51:08 2024 +0200 @@ -32,19 +32,19 @@ (* A is the true creator of X if she has sent X and X never appeared on the trace before this event. Recall that traces grow from head. *) Issues :: "[agent, agent, msg, event list] \ bool" - ("_ Issues _ with _ on _" [50, 0, 0, 50] 50) where + (\_ Issues _ with _ on _\ [50, 0, 0, 50] 50) where "(A Issues B with X on evs) = (\Y. Says A B Y \ set evs \ X \ parts {Y} \ X \ parts (spies (takeWhile (\z. z \ Says A B Y) (rev evs))))" definition (* Yields the subtrace of a given trace from its beginning to a given event *) - before :: "[event, event list] \ event list" ("before _ on _" [0, 50] 50) + before :: "[event, event list] \ event list" (\before _ on _\ [0, 50] 50) where "(before ev on evs) = takeWhile (\z. z \ ev) (rev evs)" definition (* States than an event really appears only once on a trace *) - Unique :: "[event, event list] \ bool" ("Unique _ on _" [0, 50] 50) + Unique :: "[event, event list] \ bool" (\Unique _ on _\ [0, 50] 50) where "(Unique ev on evs) = (ev \ set (tl (dropWhile (\z. z \ ev) evs)))" @@ -95,7 +95,7 @@ "expiredA T evs == authlife + T < CT evs" abbreviation - valid :: "[nat, nat] \ bool" ("valid _ wrt _" [0, 50] 50) where + valid :: "[nat, nat] \ bool" (\valid _ wrt _\ [0, 50] 50) where "valid T1 wrt T2 == T1 \ replylife + T2" (*---------------------------------------------------------------------*) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/KerberosIV_Gets.thy --- a/src/HOL/Auth/KerberosIV_Gets.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/KerberosIV_Gets.thy Fri Sep 20 19:51:08 2024 +0200 @@ -30,7 +30,7 @@ definition (* States than an event really appears only once on a trace *) - Unique :: "[event, event list] \ bool" ("Unique _ on _" [0, 50] 50) + Unique :: "[event, event list] \ bool" (\Unique _ on _\ [0, 50] 50) where "(Unique ev on evs) = (ev \ set (tl (dropWhile (\z. z \ ev) evs)))" @@ -81,7 +81,7 @@ "expiredA T evs == authlife + T < CT evs" abbreviation - valid :: "[nat, nat] \ bool" ("valid _ wrt _" [0, 50] 50) where + valid :: "[nat, nat] \ bool" (\valid _ wrt _\ [0, 50] 50) where "valid T1 wrt T2 == T1 \ replylife + T2" (*---------------------------------------------------------------------*) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/KerberosV.thy --- a/src/HOL/Auth/KerberosV.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/KerberosV.thy Fri Sep 20 19:51:08 2024 +0200 @@ -33,7 +33,7 @@ (* A is the true creator of X if she has sent X and X never appeared on the trace before this event. Recall that traces grow from head. *) Issues :: "[agent, agent, msg, event list] \ bool" - ("_ Issues _ with _ on _") where + (\_ Issues _ with _ on _\) where "A Issues B with X on evs = (\Y. Says A B Y \ set evs \ X \ parts {Y} \ X \ parts (spies (takeWhile (\z. z \ Says A B Y) (rev evs))))" @@ -86,7 +86,7 @@ "expiredA T evs == authlife + T < CT evs" abbreviation - valid :: "[nat, nat] \ bool" ("valid _ wrt _") where + valid :: "[nat, nat] \ bool" (\valid _ wrt _\) where "valid T1 wrt T2 == T1 \ replylife + T2" (*---------------------------------------------------------------------*) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/Kerberos_BAN.thy --- a/src/HOL/Auth/Kerberos_BAN.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/Kerberos_BAN.thy Fri Sep 20 19:51:08 2024 +0200 @@ -54,19 +54,19 @@ (* A is the true creator of X if she has sent X and X never appeared on the trace before this event. Recall that traces grow from head. *) Issues :: "[agent, agent, msg, event list] \ bool" - ("_ Issues _ with _ on _") where + (\_ Issues _ with _ on _\) where "A Issues B with X on evs = (\Y. Says A B Y \ set evs \ X \ parts {Y} \ X \ parts (spies (takeWhile (\z. z \ Says A B Y) (rev evs))))" definition (* Yields the subtrace of a given trace from its beginning to a given event *) - before :: "[event, event list] \ event list" ("before _ on _") + before :: "[event, event list] \ event list" (\before _ on _\) where "before ev on evs = takeWhile (\z. z \ ev) (rev evs)" definition (* States than an event really appears only once on a trace *) - Unique :: "[event, event list] \ bool" ("Unique _ on _") + Unique :: "[event, event list] \ bool" (\Unique _ on _\) where "Unique ev on evs = (ev \ set (tl (dropWhile (\z. z \ ev) evs)))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/Kerberos_BAN_Gets.thy --- a/src/HOL/Auth/Kerberos_BAN_Gets.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy Fri Sep 20 19:51:08 2024 +0200 @@ -54,12 +54,12 @@ definition (* Yields the subtrace of a given trace from its beginning to a given event *) - before :: "[event, event list] \ event list" ("before _ on _") + before :: "[event, event list] \ event list" (\before _ on _\) where "before ev on evs = takeWhile (\z. z \ ev) (rev evs)" definition (* States than an event really appears only once on a trace *) - Unique :: "[event, event list] \ bool" ("Unique _ on _") + Unique :: "[event, event list] \ bool" (\Unique _ on _\) where "Unique ev on evs = (ev \ set (tl (dropWhile (\z. z \ ev) evs)))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/Message.thy Fri Sep 20 19:51:08 2024 +0200 @@ -51,9 +51,9 @@ text\Concrete syntax: messages appear as \\A,B,NA\\, etc...\ nonterminal mtuple_args syntax - "" :: "'a \ mtuple_args" ("_") - "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" ("_,/ _") - "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" ("(2\_,/ _\)") + "" :: "'a \ mtuple_args" (\_\) + "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" (\_,/ _\) + "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" (\(2\_,/ _\)\) syntax_consts "_MTuple_args" "_MTuple" \ MPair translations @@ -61,7 +61,7 @@ "\x, y\" \ "CONST MPair x y" -definition HPair :: "[msg,msg] \ msg" ("(4Hash[_] /_)" [0, 1000]) where +definition HPair :: "[msg,msg] \ msg" (\(4Hash[_] /_)\ [0, 1000]) where \ \Message Y paired with a MAC computed with the help of X\ "Hash[X] Y == \Hash\X,Y\, Y\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/NS_Shared.thy --- a/src/HOL/Auth/NS_Shared.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/NS_Shared.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,7 +17,7 @@ (* A is the true creator of X if she has sent X and X never appeared on the trace before this event. Recall that traces grow from head. *) Issues :: "[agent, agent, msg, event list] \ bool" - ("_ Issues _ with _ on _") where + (\_ Issues _ with _ on _\) where "A Issues B with X on evs = (\Y. Says A B Y \ set evs \ X \ parts {Y} \ X \ parts (spies (takeWhile (\z. z \ Says A B Y) (rev evs))))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/Smartcard/ShoupRubin.thy --- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Fri Sep 20 19:51:08 2024 +0200 @@ -23,7 +23,7 @@ between each agent and his smartcard*) shouprubin_assumes_securemeans [iff]: "evs \ sr \ secureM" -definition Unique :: "[event, event list] => bool" ("Unique _ on _") where +definition Unique :: "[event, event list] => bool" (\Unique _ on _\) where "Unique ev on evs == ev \ set (tl (dropWhile (% z. z \ ev) evs))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/Smartcard/ShoupRubinBella.thy --- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Fri Sep 20 19:51:08 2024 +0200 @@ -30,7 +30,7 @@ between each agent and his smartcard*) shouprubin_assumes_securemeans [iff]: "evs \ srb \ secureM" -definition Unique :: "[event, event list] => bool" ("Unique _ on _") where +definition Unique :: "[event, event list] => bool" (\Unique _ on _\) where "Unique ev on evs == ev \ set (tl (dropWhile (% z. z \ ev) evs))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Fri Sep 20 19:51:08 2024 +0200 @@ -42,7 +42,7 @@ shrK_disj_pin [iff]: "shrK P \ pin Q" and crdK_disj_pin [iff]: "crdK C \ pin P" -definition legalUse :: "card => bool" ("legalUse (_)") where +definition legalUse :: "card => bool" (\legalUse (_)\) where "legalUse C == C \ stolen" primrec illegalUse :: "card => bool" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/AxCompl.thy --- a/src/HOL/Bali/AxCompl.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/AxCompl.thy Fri Sep 20 19:51:08 2024 +0200 @@ -96,7 +96,7 @@ subsubsection "init-le" definition - init_le :: "prog \ nat \ state \ bool" ("_\init\_" [51,51] 50) + init_le :: "prog \ nat \ state \ bool" (\_\init\_\ [51,51] 50) where "G\init\n = (\s. card (nyinitcls G s) \ n)" lemma init_le_def2 [simp]: "(G\init\n) s = (card (nyinitcls G s)\n)" @@ -116,7 +116,7 @@ subsubsection "Most General Triples and Formulas" definition - remember_init_state :: "state assn" ("\") + remember_init_state :: "state assn" (\\\) where "\ \ \Y s Z. s = Z" lemma remember_init_state_def2 [simp]: "\ Y = (=)" @@ -125,11 +125,11 @@ done definition - MGF ::"[state assn, term, prog] \ state triple" ("{_} _\ {_\}"[3,65,3]62) + MGF ::"[state assn, term, prog] \ state triple" (\{_} _\ {_\}\[3,65,3]62) where "{P} t\ {G\} = {P} t\ {\Y s' s. G\s \t\\ (Y,s')}" definition - MGFn :: "[nat, term, prog] \ state triple" ("{=:_} _\ {_\}"[3,65,3]62) + MGFn :: "[nat, term, prog] \ state triple" (\{=:_} _\ {_\}\[3,65,3]62) where "{=:n} t\ {G\} = {\ \. G\init\n} t\ {G\}" (* unused *) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/AxSem.thy --- a/src/HOL/Bali/AxSem.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/AxSem.thy Fri Sep 20 19:51:08 2024 +0200 @@ -48,9 +48,9 @@ Vals where "Vals x == In3 x" syntax - "_Val" :: "[pttrn] => pttrn" ("Val:_" [951] 950) - "_Var" :: "[pttrn] => pttrn" ("Var:_" [951] 950) - "_Vals" :: "[pttrn] => pttrn" ("Vals:_" [951] 950) + "_Val" :: "[pttrn] => pttrn" (\Val:_\ [951] 950) + "_Var" :: "[pttrn] => pttrn" (\Var:_\ [951] 950) + "_Vals" :: "[pttrn] => pttrn" (\Vals:_\ [951] 950) translations "\Val:v . b" == "(\v. b) \ CONST the_In1" @@ -63,7 +63,7 @@ (type) "'a assn" <= (type) "vals \ state \ 'a \ bool" definition - assn_imp :: "'a assn \ 'a assn \ bool" (infixr "\" 25) + assn_imp :: "'a assn \ 'a assn \ bool" (infixr \\\ 25) where "(P \ Q) = (\Y s Z. P Y s Z \ Q Y s Z)" lemma assn_imp_def2 [iff]: "(P \ Q) = (\Y s Z. P Y s Z \ Q Y s Z)" @@ -77,7 +77,7 @@ subsection "peek-and" definition - peek_and :: "'a assn \ (state \ bool) \ 'a assn" (infixl "\." 13) + peek_and :: "'a assn \ (state \ bool) \ 'a assn" (infixl \\.\ 13) where "(P \. p) = (\Y s Z. P Y s Z \ p s)" lemma peek_and_def2 [simp]: "peek_and P p Y s = (\Z. (P Y s Z \ p s))" @@ -117,7 +117,7 @@ subsection "assn-supd" definition - assn_supd :: "'a assn \ (state \ state) \ 'a assn" (infixl ";." 13) + assn_supd :: "'a assn \ (state \ state) \ 'a assn" (infixl \;.\ 13) where "(P ;. f) = (\Y s' Z. \s. P Y s Z \ s' = f s)" lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (\s. P Y s Z \ s' = f s)" @@ -128,7 +128,7 @@ subsection "supd-assn" definition - supd_assn :: "(state \ state) \ 'a assn \ 'a assn" (infixr ".;" 13) + supd_assn :: "(state \ state) \ 'a assn \ 'a assn" (infixr \.;\ 13) where "(f .; P) = (\Y s. P Y (f s))" @@ -148,7 +148,7 @@ subsection "subst-res" definition - subst_res :: "'a assn \ res \ 'a assn" ("_\_" [60,61] 60) + subst_res :: "'a assn \ res \ 'a assn" (\_\_\ [60,61] 60) where "P\w = (\Y. P w)" lemma subst_res_def2 [simp]: "(P\w) Y = P w" @@ -184,7 +184,7 @@ subsection "subst-Bool" definition - subst_Bool :: "'a assn \ bool \ 'a assn" ("_\=_" [60,61] 60) + subst_Bool :: "'a assn \ bool \ 'a assn" (\_\=_\ [60,61] 60) where "P\=b = (\Y s Z. \v. P (Val v) s Z \ (normal s \ the_Bool v=b))" lemma subst_Bool_def2 [simp]: @@ -204,7 +204,7 @@ where "peek_res Pf = (\Y. Pf Y Y)" syntax - "_peek_res" :: "pttrn \ 'a assn \ 'a assn" ("\_:. _" [0,3] 3) + "_peek_res" :: "pttrn \ 'a assn \ 'a assn" (\\_:. _\ [0,3] 3) syntax_consts "_peek_res" == peek_res translations @@ -231,7 +231,7 @@ subsection "ign-res" definition - ign_res :: "'a assn \ 'a assn" ("_\" [1000] 1000) + ign_res :: "'a assn \ 'a assn" (\_\\ [1000] 1000) where "P\ = (\Y s Z. \Y. P Y s Z)" lemma ign_res_def2 [simp]: "P\ Y s Z = (\Y. P Y s Z)" @@ -267,7 +267,7 @@ where "peek_st P = (\Y s. P (store s) Y s)" syntax - "_peek_st" :: "pttrn \ 'a assn \ 'a assn" ("\_.. _" [0,3] 3) + "_peek_st" :: "pttrn \ 'a assn \ 'a assn" (\\_.. _\ [0,3] 3) syntax_consts "_peek_st" == peek_st translations @@ -310,7 +310,7 @@ subsection "ign-res-eq" definition - ign_res_eq :: "'a assn \ res \ 'a assn" ("_\=_" [60,61] 60) + ign_res_eq :: "'a assn \ res \ 'a assn" (\_\=_\ [60,61] 60) where "P\=w \ (\Y:. P\ \. (\s. Y=w))" lemma ign_res_eq_def2 [simp]: "(P\=w) Y s Z = ((\Y. P Y s Z) \ Y=w)" @@ -341,7 +341,7 @@ subsection "RefVar" definition - RefVar :: "(state \ vvar \ state) \ 'a assn \ 'a assn" (infixr "..;" 13) + RefVar :: "(state \ vvar \ state) \ 'a assn \ 'a assn" (infixr \..;\ 13) where "(vf ..; P) = (\Y s. let (v,s') = vf s in P (Var v) s')" lemma RefVar_def2 [simp]: "(vf ..; P) Y s = @@ -384,34 +384,34 @@ datatype 'a triple = triple "('a assn)" "term" "('a assn)" (** should be something like triple = \'a. triple ('a assn) term ('a assn) **) - ("{(1_)}/ _\/ {(1_)}" [3,65,3] 75) + (\{(1_)}/ _\/ {(1_)}\ [3,65,3] 75) type_synonym 'a triples = "'a triple set" abbreviation var_triple :: "['a assn, var ,'a assn] \ 'a triple" - ("{(1_)}/ _=\/ {(1_)}" [3,80,3] 75) + (\{(1_)}/ _=\/ {(1_)}\ [3,80,3] 75) where "{P} e=\ {Q} == {P} In2 e\ {Q}" abbreviation expr_triple :: "['a assn, expr ,'a assn] \ 'a triple" - ("{(1_)}/ _-\/ {(1_)}" [3,80,3] 75) + (\{(1_)}/ _-\/ {(1_)}\ [3,80,3] 75) where "{P} e-\ {Q} == {P} In1l e\ {Q}" abbreviation exprs_triple :: "['a assn, expr list ,'a assn] \ 'a triple" - ("{(1_)}/ _\\/ {(1_)}" [3,65,3] 75) + (\{(1_)}/ _\\/ {(1_)}\ [3,65,3] 75) where "{P} e\\ {Q} == {P} In3 e\ {Q}" abbreviation stmt_triple :: "['a assn, stmt, 'a assn] \ 'a triple" - ("{(1_)}/ ._./ {(1_)}" [3,65,3] 75) + (\{(1_)}/ ._./ {(1_)}\ [3,65,3] 75) where "{P} .c. {Q} == {P} In1r c\ {Q}" notation (ASCII) - triple ("{(1_)}/ _>/ {(1_)}" [3,65,3]75) and - var_triple ("{(1_)}/ _=>/ {(1_)}" [3,80,3] 75) and - expr_triple ("{(1_)}/ _->/ {(1_)}" [3,80,3] 75) and - exprs_triple ("{(1_)}/ _#>/ {(1_)}" [3,65,3] 75) + triple (\{(1_)}/ _>/ {(1_)}\ [3,65,3]75) and + var_triple (\{(1_)}/ _=>/ {(1_)}\ [3,80,3] 75) and + expr_triple (\{(1_)}/ _->/ {(1_)}\ [3,80,3] 75) and + exprs_triple (\{(1_)}/ _#>/ {(1_)}\ [3,65,3] 75) lemma inj_triple: "inj (\(P,t,Q). {P} t\ {Q})" apply (rule inj_onI) @@ -423,11 +423,11 @@ done definition mtriples :: "('c \ 'sig \ 'a assn) \ ('c \ 'sig \ expr) \ - ('c \ 'sig \ 'a assn) \ ('c \ 'sig) set \ 'a triples" ("{{(1_)}/ _-\/ {(1_)} | _}"[3,65,3,65]75) where + ('c \ 'sig \ 'a assn) \ ('c \ 'sig) set \ 'a triples" (\{{(1_)}/ _-\/ {(1_)} | _}\[3,65,3,65]75) where "{{P} tf-\ {Q} | ms} = (\(C,sig). {Normal(P C sig)} tf C sig-\ {Q C sig})`ms" definition - triple_valid :: "prog \ nat \ 'a triple \ bool" ("_\_:_" [61,0, 58] 57) + triple_valid :: "prog \ nat \ 'a triple \ bool" (\_\_:_\ [61,0, 58] 57) where "G\n:t = (case t of {P} t\ {Q} \ @@ -435,23 +435,23 @@ (\Y' s'. G\s \t\\n\ (Y',s') \ Q Y' s' Z))" abbreviation - triples_valid:: "prog \ nat \ 'a triples \ bool" ("_|\_:_" [61,0, 58] 57) + triples_valid:: "prog \ nat \ 'a triples \ bool" (\_|\_:_\ [61,0, 58] 57) where "G|\n:ts == Ball ts (triple_valid G n)" notation (ASCII) - triples_valid ( "_||=_:_" [61,0, 58] 57) + triples_valid ( \_||=_:_\ [61,0, 58] 57) definition - ax_valids :: "prog \ 'b triples \ 'a triples \ bool" ("_,_|\_" [61,58,58] 57) + ax_valids :: "prog \ 'b triples \ 'a triples \ bool" (\_,_|\_\ [61,58,58] 57) where "(G,A|\ts) = (\n. G|\n:A \ G|\n:ts)" abbreviation - ax_valid :: "prog \ 'b triples \ 'a triple \ bool" ("_,_\_" [61,58,58] 57) + ax_valid :: "prog \ 'b triples \ 'a triple \ bool" (\_,_\_\ [61,58,58] 57) where "G,A \t == G,A|\{t}" notation (ASCII) - ax_valid ( "_,_|=_" [61,58,58] 57) + ax_valid ( \_,_|=_\ [61,58,58] 57) lemma triple_valid_def2: "G\n:{P} t\ {Q} = @@ -472,8 +472,8 @@ setup \map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\ inductive - ax_derivs :: "prog \ 'a triples \ 'a triples \ bool" ("_,_|\_" [61,58,58] 57) - and ax_deriv :: "prog \ 'a triples \ 'a triple \ bool" ("_,_\_" [61,58,58] 57) + ax_derivs :: "prog \ 'a triples \ 'a triples \ bool" (\_,_|\_\ [61,58,58] 57) + and ax_deriv :: "prog \ 'a triples \ 'a triple \ bool" (\_,_\_\ [61,58,58] 57) for G :: prog where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/AxSound.thy --- a/src/HOL/Bali/AxSound.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/AxSound.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,7 +10,7 @@ subsubsection "validity" definition - triple_valid2 :: "prog \ nat \ 'a triple \ bool" ("_\_\_"[61,0, 58] 57) + triple_valid2 :: "prog \ nat \ 'a triple \ bool" (\_\_\_\[61,0, 58] 57) where "G\n\t = (case t of {P} t\ {Q} \ @@ -28,7 +28,7 @@ \ definition - ax_valids2 :: "prog \ 'a triples \ 'a triples \ bool" ("_,_|\\_" [61,58,58] 57) + ax_valids2 :: "prog \ 'a triples \ 'a triples \ bool" (\_,_|\\_\ [61,58,58] 57) where "G,A|\\ts = (\n. (\t\A. G\n\t) \ (\t\ts. G\n\t))" lemma triple_valid2_def2: "G\n\{P} t\ {Q} = diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Basis.thy Fri Sep 20 19:51:08 2024 +0200 @@ -26,8 +26,8 @@ apply fast+ done -abbreviation nat3 :: nat ("3") where "3 \ Suc 2" -abbreviation nat4 :: nat ("4") where "4 \ Suc 3" +abbreviation nat3 :: nat (\3\) where "3 \ Suc 2" +abbreviation nat4 :: nat (\4\) where "4 \ Suc 3" (* irrefl_tranclI in Transitive_Closure.thy is more general *) lemma irrefl_tranclI': "r\ \ r\<^sup>+ = {} \ \x. (x, x) \ r\<^sup>+" @@ -144,7 +144,7 @@ subsubsection "sums" -notation case_sum (infixr "'(+')" 80) +notation case_sum (infixr \'(+')\ 80) primrec the_Inl :: "'a + 'b \ 'a" where "the_Inl (Inl a) = a" @@ -188,12 +188,12 @@ subsubsection "quantifiers for option type" syntax - "_Oall" :: "[pttrn, 'a option, bool] \ bool" ("(3! _:_:/ _)" [0,0,10] 10) - "_Oex" :: "[pttrn, 'a option, bool] \ bool" ("(3? _:_:/ _)" [0,0,10] 10) + "_Oall" :: "[pttrn, 'a option, bool] \ bool" (\(3! _:_:/ _)\ [0,0,10] 10) + "_Oex" :: "[pttrn, 'a option, bool] \ bool" (\(3? _:_:/ _)\ [0,0,10] 10) syntax (symbols) - "_Oall" :: "[pttrn, 'a option, bool] \ bool" ("(3\_\_:/ _)" [0,0,10] 10) - "_Oex" :: "[pttrn, 'a option, bool] \ bool" ("(3\_\_:/ _)" [0,0,10] 10) + "_Oall" :: "[pttrn, 'a option, bool] \ bool" (\(3\_\_:/ _)\ [0,0,10] 10) + "_Oex" :: "[pttrn, 'a option, bool] \ bool" (\(3\_\_:/ _)\ [0,0,10] 10) syntax_consts "_Oall" \ Ball and @@ -260,7 +260,7 @@ text \list patterns -- extends pre-defined type "pttrn" used in abstractions\ syntax - "_lpttrn" :: "[pttrn, pttrn] \ pttrn" ("_#/_" [901,900] 900) + "_lpttrn" :: "[pttrn, pttrn] \ pttrn" (\_#/_\ [901,900] 900) syntax_consts "_lpttrn" \ lsplit translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Conform.thy --- a/src/HOL/Bali/Conform.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Conform.thy Fri Sep 20 19:51:08 2024 +0200 @@ -22,7 +22,7 @@ subsubsection "extension of global store" -definition gext :: "st \ st \ bool" ("_\|_" [71,71] 70) where +definition gext :: "st \ st \ bool" (\_\|_\ [71,71] 70) where "s\|s' \ \r. \obj\globs s r: \obj'\globs s' r: tag obj'= tag obj" text \For the the proof of type soundness we will need the @@ -96,7 +96,7 @@ subsubsection "value conformance" -definition conf :: "prog \ st \ val \ ty \ bool" ("_,_\_\\_" [71,71,71,71] 70) +definition conf :: "prog \ st \ val \ ty \ bool" (\_,_\_\\_\ [71,71,71,71] 70) where "G,s\v\\T = (\T'\typeof (\a. map_option obj_ty (heap s a)) v:G\T'\T)" lemma conf_cong [simp]: "G,set_locals l s\v\\T = G,s\v\\T" @@ -178,7 +178,7 @@ subsubsection "value list conformance" definition - lconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" ("_,_\_[\\]_" [71,71,71,71] 70) + lconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" (\_,_\_[\\]_\ [71,71,71,71] 70) where "G,s\vs[\\]Ts = (\n. \T\Ts n: \v\vs n: G,s\v\\T)" lemma lconfD: "\G,s\vs[\\]Ts; Ts n = Some T\ \ G,s\(the (vs n))\\T" @@ -262,7 +262,7 @@ definition - wlconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" ("_,_\_[\\\]_" [71,71,71,71] 70) + wlconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" (\_,_\_[\\\]_\ [71,71,71,71] 70) where "G,s\vs[\\\]Ts = (\n. \T\Ts n: \ v\vs n: G,s\v\\T)" lemma wlconfD: "\G,s\vs[\\\]Ts; Ts n = Some T; vs n = Some v\ \ G,s\v\\T" @@ -341,7 +341,7 @@ subsubsection "object conformance" definition - oconf :: "prog \ st \ obj \ oref \ bool" ("_,_\_\\\_" [71,71,71,71] 70) where + oconf :: "prog \ st \ obj \ oref \ bool" (\_,_\_\\\_\ [71,71,71,71] 70) where "(G,s\obj\\\r) = (G,s\values obj[\\]var_tys G (tag obj) r \ (case r of Heap a \ is_type G (obj_ty obj) @@ -376,7 +376,7 @@ subsubsection "state conformance" definition - conforms :: "state \ env' \ bool" ("_\\_" [71,71] 70) where + conforms :: "state \ env' \ bool" (\_\\_\ [71,71] 70) where "xs\\E = (let (G, L) = E; s = snd xs; l = locals s in (\r. \obj\globs s r: G,s\obj \\\r) \ G,s\l [\\\]L \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Decl.thy --- a/src/HOL/Bali/Decl.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Decl.thy Fri Sep 20 19:51:08 2024 +0200 @@ -456,21 +456,21 @@ where "subcls1 G = {(C,D). C\Object \ (\c\class G C: super c = D)}" abbreviation - subcls1_syntax :: "prog => [qtname, qtname] => bool" ("_\_\\<^sub>C1_" [71,71,71] 70) + subcls1_syntax :: "prog => [qtname, qtname] => bool" (\_\_\\<^sub>C1_\ [71,71,71] 70) where "G\C \\<^sub>C1 D == (C,D) \ subcls1 G" abbreviation - subclseq_syntax :: "prog => [qtname, qtname] => bool" ("_\_\\<^sub>C _" [71,71,71] 70) + subclseq_syntax :: "prog => [qtname, qtname] => bool" (\_\_\\<^sub>C _\ [71,71,71] 70) where "G\C \\<^sub>C D == (C,D) \(subcls1 G)\<^sup>*" (* cf. 8.1.3 *) abbreviation - subcls_syntax :: "prog => [qtname, qtname] => bool" ("_\_\\<^sub>C _" [71,71,71] 70) + subcls_syntax :: "prog => [qtname, qtname] => bool" (\_\_\\<^sub>C _\ [71,71,71] 70) where "G\C \\<^sub>C D == (C,D) \(subcls1 G)\<^sup>+" notation (ASCII) - subcls1_syntax ("_|-_<:C1_" [71,71,71] 70) and - subclseq_syntax ("_|-_<=:C _"[71,71,71] 70) and - subcls_syntax ("_|-_<:C _"[71,71,71] 70) + subcls1_syntax (\_|-_<:C1_\ [71,71,71] 70) and + subclseq_syntax (\_|-_<=:C _\[71,71,71] 70) and + subcls_syntax (\_|-_<:C _\[71,71,71] 70) lemma subint1I: "\iface G I = Some i; J \ set (isuperIfs i)\ \ (I,J) \ subint1 G" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/DeclConcepts.thy --- a/src/HOL/Bali/DeclConcepts.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/DeclConcepts.thy Fri Sep 20 19:51:08 2024 +0200 @@ -22,8 +22,8 @@ its element type is accessible\ primrec - accessible_in :: "prog \ ty \ pname \ bool" ("_ \ _ accessible'_in _" [61,61,61] 60) and - rt_accessible_in :: "prog \ ref_ty \ pname \ bool" ("_ \ _ accessible'_in'' _" [61,61,61] 60) + accessible_in :: "prog \ ty \ pname \ bool" (\_ \ _ accessible'_in _\ [61,61,61] 60) and + rt_accessible_in :: "prog \ ref_ty \ pname \ bool" (\_ \ _ accessible'_in'' _\ [61,61,61] 60) where "G\(PrimT p) accessible_in pack = True" | accessible_in_RefT_simp: @@ -519,7 +519,7 @@ \end{itemize} \ definition - inheritable_in :: "prog \ (qtname \ memberdecl) \ pname \ bool" ("_ \ _ inheritable'_in _" [61,61,61] 60) + inheritable_in :: "prog \ (qtname \ memberdecl) \ pname \ bool" (\_ \ _ inheritable'_in _\ [61,61,61] 60) where "G\membr inheritable_in pack = (case (accmodi membr) of @@ -531,13 +531,13 @@ abbreviation Method_inheritable_in_syntax:: "prog \ (qtname \ mdecl) \ pname \ bool" - ("_ \Method _ inheritable'_in _ " [61,61,61] 60) + (\_ \Method _ inheritable'_in _ \ [61,61,61] 60) where "G\Method m inheritable_in p == G\methdMembr m inheritable_in p" abbreviation Methd_inheritable_in:: "prog \ sig \ (qtname \ methd) \ pname \ bool" - ("_ \Methd _ _ inheritable'_in _ " [61,61,61,61] 60) + (\_ \Methd _ _ inheritable'_in _ \ [61,61,61,61] 60) where "G\Methd s m inheritable_in p == G\(method s m) inheritable_in p" subsubsection "declared-in/undeclared-in" @@ -557,7 +557,7 @@ | Some c \ table_of (cfields c))" definition - declared_in :: "prog \ memberdecl \ qtname \ bool" ("_\ _ declared'_in _" [61,61,61] 60) + declared_in :: "prog \ memberdecl \ qtname \ bool" (\_\ _ declared'_in _\ [61,61,61] 60) where "G\m declared_in C = (case m of fdecl (fn,f ) \ cdeclaredfield G C fn = Some f @@ -565,12 +565,12 @@ abbreviation method_declared_in:: "prog \ (qtname \ mdecl) \ qtname \ bool" - ("_\Method _ declared'_in _" [61,61,61] 60) + (\_\Method _ declared'_in _\ [61,61,61] 60) where "G\Method m declared_in C == G\mdecl (mthd m) declared_in C" abbreviation methd_declared_in:: "prog \ sig \(qtname \ methd) \ qtname \ bool" - ("_\Methd _ _ declared'_in _" [61,61,61,61] 60) + (\_\Methd _ _ declared'_in _\ [61,61,61,61] 60) where "G\Methd s m declared_in C == G\mdecl (s,mthd m) declared_in C" lemma declared_in_classD: @@ -579,7 +579,7 @@ (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def) definition - undeclared_in :: "prog \ memberid \ qtname \ bool" ("_\ _ undeclared'_in _" [61,61,61] 60) + undeclared_in :: "prog \ memberid \ qtname \ bool" (\_\ _ undeclared'_in _\ [61,61,61] 60) where "G\m undeclared_in C = (case m of fid fn \ cdeclaredfield G C fn = None @@ -595,7 +595,7 @@ inductive members :: "prog \ (qtname \ memberdecl) \ qtname \ bool" - ("_ \ _ member'_of _" [61,61,61] 60) + (\_ \ _ member'_of _\ [61,61,61] 60) for G :: prog where @@ -614,21 +614,21 @@ abbreviation method_member_of:: "prog \ (qtname \ mdecl) \ qtname \ bool" - ("_ \Method _ member'_of _" [61,61,61] 60) + (\_ \Method _ member'_of _\ [61,61,61] 60) where "G\Method m member_of C == G\(methdMembr m) member_of C" abbreviation methd_member_of:: "prog \ sig \ (qtname \ methd) \ qtname \ bool" - ("_ \Methd _ _ member'_of _" [61,61,61,61] 60) + (\_ \Methd _ _ member'_of _\ [61,61,61,61] 60) where "G\Methd s m member_of C == G\(method s m) member_of C" abbreviation fieldm_member_of:: "prog \ vname \ (qtname \ field) \ qtname \ bool" - ("_ \Field _ _ member'_of _" [61,61,61] 60) + (\_ \Field _ _ member'_of _\ [61,61,61] 60) where "G\Field n f member_of C == G\fieldm n f member_of C" definition - inherits :: "prog \ qtname \ (qtname \ memberdecl) \ bool" ("_ \ _ inherits _" [61,61,61] 60) + inherits :: "prog \ qtname \ (qtname \ memberdecl) \ bool" (\_ \ _ inherits _\ [61,61,61] 60) where "G\C inherits m = (G\m inheritable_in (pid C) \ G\memberid m undeclared_in C \ @@ -639,7 +639,7 @@ definition - member_in :: "prog \ (qtname \ memberdecl) \ qtname \ bool" ("_ \ _ member'_in _" [61,61,61] 60) + member_in :: "prog \ (qtname \ memberdecl) \ qtname \ bool" (\_ \ _ member'_in _\ [61,61,61] 60) where "G\m member_in C = (\ provC. G\ C \\<^sub>C provC \ G \ m member_of provC)" text \A member is in a class if it is member of the class or a superclass. If a member is in a class we can select this member. This additional notion @@ -648,12 +648,12 @@ abbreviation method_member_in:: "prog \ (qtname \ mdecl) \ qtname \ bool" - ("_ \Method _ member'_in _" [61,61,61] 60) + (\_ \Method _ member'_in _\ [61,61,61] 60) where "G\Method m member_in C == G\(methdMembr m) member_in C" abbreviation methd_member_in:: "prog \ sig \ (qtname \ methd) \ qtname \ bool" - ("_ \Methd _ _ member'_in _" [61,61,61,61] 60) + (\_ \Methd _ _ member'_in _\ [61,61,61,61] 60) where "G\Methd s m member_in C == G\(method s m) member_in C" lemma member_inD: "G\m member_in C @@ -678,7 +678,7 @@ inductive stat_overridesR :: "prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" - ("_ \ _ overrides\<^sub>S _" [61,61,61] 60) + (\_ \ _ overrides\<^sub>S _\ [61,61,61] 60) for G :: prog where @@ -697,7 +697,7 @@ inductive overridesR :: "prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" - ("_ \ _ overrides _" [61,61,61] 60) + (\_ \ _ overrides _\ [61,61,61] 60) for G :: prog where @@ -716,18 +716,18 @@ abbreviation (input) sig_stat_overrides:: "prog \ sig \ (qtname \ methd) \ (qtname \ methd) \ bool" - ("_,_\ _ overrides\<^sub>S _" [61,61,61,61] 60) + (\_,_\ _ overrides\<^sub>S _\ [61,61,61,61] 60) where "G,s\new overrides\<^sub>S old == G\(qmdecl s new) overrides\<^sub>S (qmdecl s old)" abbreviation (input) sig_overrides:: "prog \ sig \ (qtname \ methd) \ (qtname \ methd) \ bool" - ("_,_\ _ overrides _" [61,61,61,61] 60) + (\_,_\ _ overrides _\ [61,61,61,61] 60) where "G,s\new overrides old == G\(qmdecl s new) overrides (qmdecl s old)" subsubsection "Hiding" definition - hides :: "prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" ("_\ _ hides _" [61,61,61] 60) + hides :: "prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" (\_\ _ hides _\ [61,61,61] 60) where "G\new hides old = (is_static new \ msig new = msig old \ @@ -738,7 +738,7 @@ abbreviation sig_hides:: "prog \ sig \ (qtname \ methd) \ (qtname \ methd) \ bool" - ("_,_\ _ hides _" [61,61,61,61] 60) + (\_,_\ _ hides _\ [61,61,61,61] 60) where "G,s\new hides old == G\(qmdecl s new) hides (qmdecl s old)" lemma hidesI: @@ -788,7 +788,7 @@ subsubsection "permits access" definition - permits_acc :: "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" ("_ \ _ in _ permits'_acc'_from _" [61,61,61,61] 60) + permits_acc :: "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" (\_ \ _ in _ permits'_acc'_from _\ [61,61,61,61] 60) where "G\membr in cls permits_acc_from accclass = (case (accmodi membr) of @@ -824,9 +824,9 @@ inductive accessible_fromR :: "prog \ qtname \ (qtname \ memberdecl) \ qtname \ bool" and accessible_from :: "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" - ("_ \ _ of _ accessible'_from _" [61,61,61,61] 60) + (\_ \ _ of _ accessible'_from _\ [61,61,61,61] 60) and method_accessible_from :: "prog \ (qtname \ mdecl) \ qtname \ qtname \ bool" - ("_ \Method _ of _ accessible'_from _" [61,61,61,61] 60) + (\_ \Method _ of _ accessible'_from _\ [61,61,61,61] 60) for G :: prog and accclass :: qtname where "G\membr of cls accessible_from accclass \ accessible_fromR G accclass membr cls" @@ -851,7 +851,7 @@ abbreviation methd_accessible_from:: "prog \ sig \ (qtname \ methd) \ qtname \ qtname \ bool" - ("_ \Methd _ _ of _ accessible'_from _" [61,61,61,61,61] 60) + (\_ \Methd _ _ of _ accessible'_from _\ [61,61,61,61,61] 60) where "G\Methd s m of cls accessible_from accclass == G\(method s m) of cls accessible_from accclass" @@ -859,7 +859,7 @@ abbreviation field_accessible_from:: "prog \ vname \ (qtname \ field) \ qtname \ qtname \ bool" - ("_ \Field _ _ of _ accessible'_from _" [61,61,61,61,61] 60) + (\_ \Field _ _ of _ accessible'_from _\ [61,61,61,61,61] 60) where "G\Field fn f of C accessible_from accclass == G\(fieldm fn f) of C accessible_from accclass" @@ -867,9 +867,9 @@ inductive dyn_accessible_fromR :: "prog \ qtname \ (qtname \ memberdecl) \ qtname \ bool" and dyn_accessible_from' :: "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" - ("_ \ _ in _ dyn'_accessible'_from _" [61,61,61,61] 60) + (\_ \ _ in _ dyn'_accessible'_from _\ [61,61,61,61] 60) and method_dyn_accessible_from :: "prog \ (qtname \ mdecl) \ qtname \ qtname \ bool" - ("_ \Method _ in _ dyn'_accessible'_from _" [61,61,61,61] 60) + (\_ \Method _ in _ dyn'_accessible'_from _\ [61,61,61,61] 60) for G :: prog and accclass :: qtname where "G\membr in C dyn_accessible_from accC \ dyn_accessible_fromR G accC membr C" @@ -890,7 +890,7 @@ abbreviation methd_dyn_accessible_from:: "prog \ sig \ (qtname \ methd) \ qtname \ qtname \ bool" - ("_ \Methd _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60) + (\_ \Methd _ _ in _ dyn'_accessible'_from _\ [61,61,61,61,61] 60) where "G\Methd s m in C dyn_accessible_from accC == G\(method s m) in C dyn_accessible_from accC" @@ -898,7 +898,7 @@ abbreviation field_dyn_accessible_from:: "prog \ vname \ (qtname \ field) \ qtname \ qtname \ bool" - ("_ \Field _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60) + (\_ \Field _ _ in _ dyn'_accessible'_from _\ [61,61,61,61,61] 60) where "G\Field fn f in dynC dyn_accessible_from accC == G\(fieldm fn f) in dynC dyn_accessible_from accC" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/DefiniteAssignment.thy --- a/src/HOL/Bali/DefiniteAssignment.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/DefiniteAssignment.thy Fri Sep 20 19:51:08 2024 +0200 @@ -428,15 +428,15 @@ subsection \Lifting set operations to range of tables (map to a set)\ definition - union_ts :: "('a,'b) tables \ ('a,'b) tables \ ('a,'b) tables" ("_ \\ _" [67,67] 65) + union_ts :: "('a,'b) tables \ ('a,'b) tables \ ('a,'b) tables" (\_ \\ _\ [67,67] 65) where "A \\ B = (\ k. A k \ B k)" definition - intersect_ts :: "('a,'b) tables \ ('a,'b) tables \ ('a,'b) tables" ("_ \\ _" [72,72] 71) + intersect_ts :: "('a,'b) tables \ ('a,'b) tables \ ('a,'b) tables" (\_ \\ _\ [72,72] 71) where "A \\ B = (\k. A k \ B k)" definition - all_union_ts :: "('a,'b) tables \ 'b set \ ('a,'b) tables" (infixl "\\\<^sub>\" 40) + all_union_ts :: "('a,'b) tables \ 'b set \ ('a,'b) tables" (infixl \\\\<^sub>\\ 40) where "(A \\\<^sub>\ B) = (\ k. A k \ B)" subsubsection \Binary union of tables\ @@ -519,7 +519,7 @@ *) definition - range_inter_ts :: "('a,'b) tables \ 'b set" ("\\_" 80) + range_inter_ts :: "('a,'b) tables \ 'b set" (\\\_\ 80) where "\\A = {x |x. \ k. x \ A k}" text \ @@ -532,7 +532,7 @@ \ inductive - da :: "env \ lname set \ term \ assigned \ bool" ("_\ _ \_\ _" [65,65,65,65] 71) + da :: "env \ lname set \ term \ assigned \ bool" (\_\ _ \_\ _\ [65,65,65,65] 71) where Skip: "Env\ B \\Skip\\ \nrm=B,brk=\ l. UNIV\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Eval.thy --- a/src/HOL/Bali/Eval.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Eval.thy Fri Sep 20 19:51:08 2024 +0200 @@ -125,19 +125,19 @@ \ abbreviation - dummy_res :: "vals" ("\") + dummy_res :: "vals" (\\\) where "\ == In1 Unit" abbreviation (input) - val_inj_vals ("\_\\<^sub>e" 1000) + val_inj_vals (\\_\\<^sub>e\ 1000) where "\e\\<^sub>e == In1 e" abbreviation (input) - var_inj_vals ("\_\\<^sub>v" 1000) + var_inj_vals (\\_\\<^sub>v\ 1000) where "\v\\<^sub>v == In2 v" abbreviation (input) - lst_inj_vals ("\_\\<^sub>l" 1000) + lst_inj_vals (\\_\\<^sub>l\ 1000) where "\es\\<^sub>l == In3 es" definition undefined3 :: "('al + 'ar, 'b, 'c) sum3 \ vals" where @@ -170,7 +170,7 @@ done definition - fits :: "prog \ st \ val \ ty \ bool" ("_,_\_ fits _"[61,61,61,61]60) + fits :: "prog \ st \ val \ ty \ bool" (\_,_\_ fits _\[61,61,61,61]60) where "G,s\a' fits T = ((\rt. T=RefT rt) \ a'=Null \ G\obj_ty(lookup_obj s a')\T)" lemma fits_Null [simp]: "G,s\Null fits T" @@ -194,7 +194,7 @@ done definition - catch :: "prog \ state \ qtname \ bool" ("_,_\catch _"[61,61,61]60) where + catch :: "prog \ state \ qtname \ bool" (\_,_\catch _\[61,61,61]60) where "G,s\catch C = (\xc. abrupt s=Some (Xcpt xc) \ G,store s\Addr (the_Loc xc) fits Class C)" @@ -469,7 +469,7 @@ subsubsection "evaluation judgments" inductive - halloc :: "[prog,state,obj_tag,loc,state]\bool" ("_\_ \halloc _\_\ _"[61,61,61,61,61]60) for G::prog + halloc :: "[prog,state,obj_tag,loc,state]\bool" (\_\_ \halloc _\_\ _\[61,61,61,61,61]60) for G::prog where \ \allocating objects on the heap, cf. 12.5\ Abrupt: @@ -481,7 +481,7 @@ \ G\Norm s \halloc oi\a\ (x,init_obj G oi' (Heap a) s)" -inductive sxalloc :: "[prog,state,state]\bool" ("_\_ \sxalloc\ _"[61,61,61]60) for G::prog +inductive sxalloc :: "[prog,state,state]\bool" (\_\_ \sxalloc\ _\[61,61,61]60) for G::prog where \ \allocating exception objects for standard exceptions (other than OutOfMemory)\ @@ -498,12 +498,12 @@ inductive - eval :: "[prog,state,term,vals,state]\bool" ("_\_ \_\\ '(_, _')" [61,61,80,0,0]60) - and exec ::"[prog,state,stmt ,state]\bool"("_\_ \_\ _" [61,61,65, 61]60) - and evar ::"[prog,state,var ,vvar,state]\bool"("_\_ \_=\_\ _"[61,61,90,61,61]60) - and eval'::"[prog,state,expr ,val ,state]\bool"("_\_ \_-\_\ _"[61,61,80,61,61]60) + eval :: "[prog,state,term,vals,state]\bool" (\_\_ \_\\ '(_, _')\ [61,61,80,0,0]60) + and exec ::"[prog,state,stmt ,state]\bool"(\_\_ \_\ _\ [61,61,65, 61]60) + and evar ::"[prog,state,var ,vvar,state]\bool"(\_\_ \_=\_\ _\[61,61,90,61,61]60) + and eval'::"[prog,state,expr ,val ,state]\bool"(\_\_ \_-\_\ _\[61,61,80,61,61]60) and evals::"[prog,state,expr list , - val list ,state]\bool"("_\_ \_\\_\ _"[61,61,61,61,61]60) + val list ,state]\bool"(\_\_ \_\\_\ _\[61,61,61,61,61]60) for G::prog where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Evaln.thy --- a/src/HOL/Bali/Evaln.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Evaln.thy Fri Sep 20 19:51:08 2024 +0200 @@ -29,15 +29,15 @@ inductive evaln :: "[prog, state, term, nat, vals, state] \ bool" - ("_\_ \_\\_\ '(_, _')" [61,61,80,61,0,0] 60) + (\_\_ \_\\_\ '(_, _')\ [61,61,80,61,0,0] 60) and evarn :: "[prog, state, var, vvar, nat, state] \ bool" - ("_\_ \_=\_\_\ _" [61,61,90,61,61,61] 60) + (\_\_ \_=\_\_\ _\ [61,61,90,61,61,61] 60) and eval_n:: "[prog, state, expr, val, nat, state] \ bool" - ("_\_ \_-\_\_\ _" [61,61,80,61,61,61] 60) + (\_\_ \_-\_\_\ _\ [61,61,80,61,61,61] 60) and evalsn :: "[prog, state, expr list, val list, nat, state] \ bool" - ("_\_ \_\\_\_\ _" [61,61,61,61,61,61] 60) + (\_\_ \_\\_\_\ _\ [61,61,61,61,61,61] 60) and execn :: "[prog, state, stmt, nat, state] \ bool" - ("_\_ \_\_\ _" [61,61,65, 61,61] 60) + (\_\_ \_\_\ _\ [61,61,65, 61,61] 60) for G :: prog where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/State.thy --- a/src/HOL/Bali/State.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/State.thy Fri Sep 20 19:51:08 2024 +0200 @@ -183,7 +183,7 @@ done definition - in_bounds :: "int \ int \ bool" ("(_/ in'_bounds _)" [50, 51] 50) + in_bounds :: "int \ int \ bool" (\(_/ in'_bounds _)\ [50, 51] 50) where "i in_bounds k = (0 \ i \ i < k)" definition @@ -306,11 +306,11 @@ subsection "update" definition - gupd :: "oref \ obj \ st \ st" ("gupd'(_\_')" [10, 10] 1000) + gupd :: "oref \ obj \ st \ st" (\gupd'(_\_')\ [10, 10] 1000) where "gupd r obj = case_st (\g l. st (g(r\obj)) l)" definition - lupd :: "lname \ val \ st \ st" ("lupd'(_\_')" [10, 10] 1000) + lupd :: "lname \ val \ st \ st" (\lupd'(_\_')\ [10, 10] 1000) where "lupd vn v = case_st (\g l. st g (l(vn\v)))" definition diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Table.thy --- a/src/HOL/Bali/Table.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Table.thy Fri Sep 20 19:51:08 2024 +0200 @@ -267,25 +267,25 @@ where "Un_tables ts = (\k. \t\ts. t k)" definition overrides_t :: "('a, 'b) tables \ ('a, 'b) tables \ ('a, 'b) tables" - (infixl "\\" 100) + (infixl \\\\ 100) where "s \\ t = (\k. if t k = {} then s k else t k)" definition hidings_entails :: "('a, 'b) tables \ ('a, 'c) tables \ ('b \ 'c \ bool) \ bool" - ("_ hidings _ entails _" 20) + (\_ hidings _ entails _\ 20) where "(t hidings s entails R) = (\k. \x\t k. \y\s k. R x y)" definition \ \variant for unique table:\ hiding_entails :: "('a, 'b) table \ ('a, 'c) table \ ('b \ 'c \ bool) \ bool" - ("_ hiding _ entails _" 20) + (\_ hiding _ entails _\ 20) where "(t hiding s entails R) = (\k. \x\t k: \y\s k: R x y)" definition \ \variant for a unique table and conditional overriding:\ cond_hiding_entails :: "('a, 'b) table \ ('a, 'c) table \ ('b \ 'c \ bool) \ ('b \ 'c \ bool) \ bool" - ("_ hiding _ under _ entails _" 20) + (\_ hiding _ under _ entails _\ 20) where "(t hiding s under C entails R) = (\k. \x\t k: \y\s k: C x y \ R x y)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Term.thy --- a/src/HOL/Bali/Term.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Term.thy Fri Sep 20 19:51:08 2024 +0200 @@ -142,7 +142,7 @@ datatype var = LVar lname \ \local variable (incl. parameters)\ - | FVar qtname qtname bool expr vname ("{_,_,_}_.._"[10,10,10,85,99]90) + | FVar qtname qtname bool expr vname (\{_,_,_}_.._\[10,10,10,85,99]90) \ \class field\ \ \\<^term>\{accC,statDeclC,stat}e..fn\\ \ \\accC\: accessing class (static class were\ @@ -152,7 +152,7 @@ \ \\stat\: static or instance field?\ \ \\e\: reference to object\ \ \\fn\: field name\ - | AVar expr expr ("_.[_]"[90,10 ]90) + | AVar expr expr (\_.[_]\[90,10 ]90) \ \array component\ \ \\<^term>\e1.[e2]\: e1 array reference; e2 index\ | InsInitV stmt var @@ -161,10 +161,10 @@ and expr = NewC qtname \ \class instance creation\ - | NewA ty expr ("New _[_]"[99,10 ]85) + | NewA ty expr (\New _[_]\[99,10 ]85) \ \array creation\ | Cast ty expr \ \type cast\ - | Inst expr ref_ty ("_ InstOf _"[85,99] 85) + | Inst expr ref_ty (\_ InstOf _\[85,99] 85) \ \instanceof\ | Lit val \ \literal value, references not allowed\ | UnOp unop expr \ \unary operation\ @@ -172,11 +172,11 @@ | Super \ \special Super keyword\ | Acc var \ \variable access\ - | Ass var expr ("_:=_" [90,85 ]85) + | Ass var expr (\_:=_\ [90,85 ]85) \ \variable assign\ - | Cond expr expr expr ("_ ? _ : _" [85,85,80]80) \ \conditional\ + | Cond expr expr expr (\_ ? _ : _\ [85,85,80]80) \ \conditional\ | Call qtname ref_ty inv_mode expr mname "(ty list)" "(expr list)" - ("{_,_,_}_\_'( {_}_')"[10,10,10,85,99,10,10]85) + (\{_,_,_}_\_'( {_}_')\[10,10,10,85,99,10,10]85) \ \method call\ \ \\<^term>\{accC,statT,mode}e\mn({pTs}args)\ "\ \ \\accC\: accessing class (static class were\ @@ -199,20 +199,20 @@ and stmt = Skip \ \empty statement\ | Expr expr \ \expression statement\ - | Lab jump stmt ("_\ _" [ 99,66]66) + | Lab jump stmt (\_\ _\ [ 99,66]66) \ \labeled statement; handles break\ - | Comp stmt stmt ("_;; _" [ 66,65]65) - | If' expr stmt stmt ("If'(_') _ Else _" [ 80,79,79]70) - | Loop label expr stmt ("_\ While'(_') _" [ 99,80,79]70) + | Comp stmt stmt (\_;; _\ [ 66,65]65) + | If' expr stmt stmt (\If'(_') _ Else _\ [ 80,79,79]70) + | Loop label expr stmt (\_\ While'(_') _\ [ 99,80,79]70) | Jmp jump \ \break, continue, return\ | Throw expr - | TryC stmt qtname vname stmt ("Try _ Catch'(_ _') _" [79,99,80,79]70) + | TryC stmt qtname vname stmt (\Try _ Catch'(_ _') _\ [79,99,80,79]70) \ \\<^term>\Try c1 Catch(C vn) c2\\ \ \\c1\: block were exception may be thrown\ \ \\C\: execption class to catch\ \ \\vn\: local name for exception used in \c2\\ \ \\c2\: block to execute when exception is cateched\ - | Fin stmt stmt ("_ Finally _" [ 79,79]70) + | Fin stmt stmt (\_ Finally _\ [ 79,79]70) | FinA abopt stmt \ \Save abruption of first statement\ \ \technical term for smallstep sem.)\ | Init qtname \ \class initialization\ @@ -245,11 +245,11 @@ abbreviation this :: expr where "this == Acc (LVar This)" -abbreviation LAcc :: "vname \ expr" ("!!") +abbreviation LAcc :: "vname \ expr" (\!!\) where "!!v == Acc (LVar (EName (VNam v)))" abbreviation - LAss :: "vname \ expr \stmt" ("_:==_" [90,85] 85) + LAss :: "vname \ expr \stmt" (\_:==_\ [90,85] 85) where "v:==e == Expr (Ass (LVar (EName (VNam v))) e)" abbreviation @@ -274,19 +274,19 @@ \ abbreviation (input) - expr_inj_term :: "expr \ term" ("\_\\<^sub>e" 1000) + expr_inj_term :: "expr \ term" (\\_\\<^sub>e\ 1000) where "\e\\<^sub>e == In1l e" abbreviation (input) - stmt_inj_term :: "stmt \ term" ("\_\\<^sub>s" 1000) + stmt_inj_term :: "stmt \ term" (\\_\\<^sub>s\ 1000) where "\c\\<^sub>s == In1r c" abbreviation (input) - var_inj_term :: "var \ term" ("\_\\<^sub>v" 1000) + var_inj_term :: "var \ term" (\\_\\<^sub>v\ 1000) where "\v\\<^sub>v == In2 v" abbreviation (input) - lst_inj_term :: "expr list \ term" ("\_\\<^sub>l" 1000) + lst_inj_term :: "expr list \ term" (\\_\\<^sub>l\ 1000) where "\es\\<^sub>l == In3 es" text \It seems to be more elegant to have an overloaded injection like the @@ -294,7 +294,7 @@ \ class inj_term = - fixes inj_term:: "'a \ term" ("\_\" 1000) + fixes inj_term:: "'a \ term" (\\_\\ 1000) text \How this overloaded injections work can be seen in the theory \DefiniteAssignment\. Other big inductive relations on diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Trans.thy --- a/src/HOL/Bali/Trans.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Trans.thy Fri Sep 20 19:51:08 2024 +0200 @@ -56,7 +56,7 @@ where "SKIP == Lit Unit" inductive - step :: "[prog,term \ state,term \ state] \ bool" ("_\_ \1 _"[61,82,82] 81) + step :: "[prog,term \ state,term \ state] \ bool" (\_\_ \1 _\[61,82,82] 81) for G :: prog where @@ -344,11 +344,11 @@ "G\(\InsInitE Skip SKIP\,Norm s) \1 (\SKIP\,Norm s)" abbreviation - stepn:: "[prog, term \ state,nat,term \ state] \ bool" ("_\_ \_ _"[61,82,82] 81) + stepn:: "[prog, term \ state,nat,term \ state] \ bool" (\_\_ \_ _\[61,82,82] 81) where "G\p \n p' \ (p,p') \ {(x, y). step G x y}^^n" abbreviation - steptr:: "[prog,term \ state,term \ state] \ bool" ("_\_ \* _"[61,82,82] 81) + steptr:: "[prog,term \ state,term \ state] \ bool" (\_\_ \* _\[61,82,82] 81) where "G\p \* p' \ (p,p') \ {(x, y). step G x y}\<^sup>*" (* Equivalenzen: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/Type.thy --- a/src/HOL/Bali/Type.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/Type.thy Fri Sep 20 19:51:08 2024 +0200 @@ -33,7 +33,7 @@ abbreviation "NT == RefT NullT" abbreviation "Iface I == RefT (IfaceT I)" abbreviation "Class C == RefT (ClassT C)" -abbreviation Array :: "ty \ ty" ("_.[]" [90] 90) +abbreviation Array :: "ty \ ty" (\_.[]\ [90] 90) where "T.[] == RefT (ArrayT T)" definition diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/TypeRel.thy --- a/src/HOL/Bali/TypeRel.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/TypeRel.thy Fri Sep 20 19:51:08 2024 +0200 @@ -38,21 +38,21 @@ abbreviation - subint1_syntax :: "prog => [qtname, qtname] => bool" ("_\_\I1_" [71,71,71] 70) + subint1_syntax :: "prog => [qtname, qtname] => bool" (\_\_\I1_\ [71,71,71] 70) where "G\I \I1 J == (I,J) \ subint1 G" abbreviation - subint_syntax :: "prog => [qtname, qtname] => bool" ("_\_\I _" [71,71,71] 70) + subint_syntax :: "prog => [qtname, qtname] => bool" (\_\_\I _\ [71,71,71] 70) where "G\I \I J == (I,J) \(subint1 G)\<^sup>*" \ \cf. 9.1.3\ abbreviation - implmt1_syntax :: "prog => [qtname, qtname] => bool" ("_\_\1_" [71,71,71] 70) + implmt1_syntax :: "prog => [qtname, qtname] => bool" (\_\_\1_\ [71,71,71] 70) where "G\C \1 I == (C,I) \ implmt1 G" notation (ASCII) - subint1_syntax ("_|-_<:I1_" [71,71,71] 70) and - subint_syntax ("_|-_<=:I _"[71,71,71] 70) and - implmt1_syntax ("_|-_~>1_" [71,71,71] 70) + subint1_syntax (\_|-_<:I1_\ [71,71,71] 70) and + subint_syntax (\_|-_<=:I _\[71,71,71] 70) and + implmt1_syntax (\_|-_~>1_\ [71,71,71] 70) subsubsection "subclass and subinterface relations" @@ -335,7 +335,7 @@ done inductive \ \implementation, cf. 8.1.4\ - implmt :: "prog \ qtname \ qtname \ bool" ("_\_\_" [71,71,71] 70) + implmt :: "prog \ qtname \ qtname \ bool" (\_\_\_\ [71,71,71] 70) for G :: prog where direct: "G\C\1J \ G\C\J" @@ -371,7 +371,7 @@ inductive \ \widening, viz. method invocation conversion, cf. 5.3 i.e. kind of syntactic subtyping\ - widen :: "prog \ ty \ ty \ bool" ("_\_\_" [71,71,71] 70) + widen :: "prog \ ty \ ty \ bool" (\_\_\_\ [71,71,71] 70) for G :: prog where refl: "G\T\T" \ \identity conversion, cf. 5.1.1\ @@ -569,7 +569,7 @@ done definition - widens :: "prog \ [ty list, ty list] \ bool" ("_\_[\]_" [71,71,71] 70) + widens :: "prog \ [ty list, ty list] \ bool" (\_\_[\]_\ [71,71,71] 70) where "G\Ts[\]Ts' = list_all2 (\T T'. G\T\T') Ts Ts'" lemma widens_Nil [simp]: "G\[][\][]" @@ -595,7 +595,7 @@ (* more detailed than necessary for type-safety, see above rules. *) inductive \ \narrowing reference conversion, cf. 5.1.5\ - narrow :: "prog \ ty \ ty \ bool" ("_\_\_" [71,71,71] 70) + narrow :: "prog \ ty \ ty \ bool" (\_\_\_\ [71,71,71] 70) for G :: prog where subcls: "G\C\\<^sub>C D \ G\ Class D\Class C" @@ -646,7 +646,7 @@ subsubsection "casting relation" inductive \ \casting conversion, cf. 5.5\ - cast :: "prog \ ty \ ty \ bool" ("_\_\? _" [71,71,71] 70) + cast :: "prog \ ty \ ty \ bool" (\_\_\? _\ [71,71,71] 70) for G :: prog where widen: "G\S\T \ G\S\? T" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/TypeSafe.thy --- a/src/HOL/Bali/TypeSafe.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/TypeSafe.thy Fri Sep 20 19:51:08 2024 +0200 @@ -93,7 +93,7 @@ subsubsection "result conformance" definition - assign_conforms :: "st \ (val \ state \ state) \ ty \ env' \ bool" ("_\|_\_\\_" [71,71,71,71] 70) + assign_conforms :: "st \ (val \ state \ state) \ ty \ env' \ bool" (\_\|_\_\\_\ [71,71,71,71] 70) where "s\|f\T\\E = ((\s' w. Norm s'\\E \ fst E,s'\w\\T \ s\|s' \ assign f w (Norm s')\\E) \ @@ -101,7 +101,7 @@ definition - rconf :: "prog \ lenv \ st \ term \ vals \ tys \ bool" ("_,_,_\_\_\\_" [71,71,71,71,71,71] 70) + rconf :: "prog \ lenv \ st \ term \ vals \ tys \ bool" (\_,_,_\_\_\\_\ [71,71,71,71,71,71] 70) where "G,L,s\t\v\\T = (case T of @@ -328,7 +328,7 @@ declare fun_upd_apply [simp del] definition - DynT_prop :: "[prog,inv_mode,qtname,ref_ty] \ bool" ("_\_\_\_"[71,71,71,71]70) + DynT_prop :: "[prog,inv_mode,qtname,ref_ty] \ bool" (\_\_\_\_\[71,71,71,71]70) where "G\mode\D\t = (mode = IntVir \ is_class G D \ (if (\T. t=ArrayT T) then D=Object else G\Class D\RefT t))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/WellForm.thy --- a/src/HOL/Bali/WellForm.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/WellForm.thy Fri Sep 20 19:51:08 2024 +0200 @@ -323,7 +323,7 @@ \ (* to Table *) definition - entails :: "('a,'b) table \ ('b \ bool) \ bool" ("_ entails _" 20) + entails :: "('a,'b) table \ ('b \ bool) \ bool" (\_ entails _\ 20) where "(t entails P) = (\k. \ x \ t k: P x)" lemma entailsD: @@ -2093,7 +2093,7 @@ Array Object \ primrec valid_lookup_cls:: "prog \ ref_ty \ qtname \ bool \ bool" - ("_,_ \ _ valid'_lookup'_cls'_for _" [61,61,61,61] 60) + (\_,_ \ _ valid'_lookup'_cls'_for _\ [61,61,61,61] 60) where "G,NullT \ dynC valid_lookup_cls_for static_membr = False" | "G,IfaceT I \ dynC valid_lookup_cls_for static_membr diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Bali/WellType.thy --- a/src/HOL/Bali/WellType.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Bali/WellType.thy Fri Sep 20 19:51:08 2024 +0200 @@ -249,12 +249,12 @@ (type) "tys" <= (type) "ty + ty list" -inductive wt :: "env \ dyn_ty \ [term,tys] \ bool" ("_,_\_\_" [51,51,51,51] 50) - and wt_stmt :: "env \ dyn_ty \ stmt \ bool" ("_,_\_\\" [51,51,51] 50) - and ty_expr :: "env \ dyn_ty \ [expr ,ty ] \ bool" ("_,_\_\-_" [51,51,51,51] 50) - and ty_var :: "env \ dyn_ty \ [var ,ty ] \ bool" ("_,_\_\=_" [51,51,51,51] 50) +inductive wt :: "env \ dyn_ty \ [term,tys] \ bool" (\_,_\_\_\ [51,51,51,51] 50) + and wt_stmt :: "env \ dyn_ty \ stmt \ bool" (\_,_\_\\\ [51,51,51] 50) + and ty_expr :: "env \ dyn_ty \ [expr ,ty ] \ bool" (\_,_\_\-_\ [51,51,51,51] 50) + and ty_var :: "env \ dyn_ty \ [var ,ty ] \ bool" (\_,_\_\=_\ [51,51,51,51] 50) and ty_exprs :: "env \ dyn_ty \ [expr list, ty list] \ bool" - ("_,_\_\\_" [51,51,51,51] 50) + (\_,_\_\\_\ [51,51,51,51] 50) where "E,dt\s\\ \ E,dt\In1r s\Inl (PrimT Void)" @@ -426,31 +426,31 @@ (* for purely static typing *) abbreviation - wt_syntax :: "env \ [term,tys] \ bool" ("_\_\_" [51,51,51] 50) + wt_syntax :: "env \ [term,tys] \ bool" (\_\_\_\ [51,51,51] 50) where "E\t\T == E,empty_dt\t\ T" abbreviation - wt_stmt_syntax :: "env \ stmt \ bool" ("_\_\\" [51,51 ] 50) + wt_stmt_syntax :: "env \ stmt \ bool" (\_\_\\\ [51,51 ] 50) where "E\s\\ == E\In1r s \ Inl (PrimT Void)" abbreviation - ty_expr_syntax :: "env \ [expr, ty] \ bool" ("_\_\-_" [51,51,51] 50) + ty_expr_syntax :: "env \ [expr, ty] \ bool" (\_\_\-_\ [51,51,51] 50) where "E\e\-T == E\In1l e \ Inl T" abbreviation - ty_var_syntax :: "env \ [var, ty] \ bool" ("_\_\=_" [51,51,51] 50) + ty_var_syntax :: "env \ [var, ty] \ bool" (\_\_\=_\ [51,51,51] 50) where "E\e\=T == E\In2 e \ Inl T" abbreviation - ty_exprs_syntax :: "env \ [expr list, ty list] \ bool" ("_\_\\_" [51,51,51] 50) + ty_exprs_syntax :: "env \ [expr list, ty list] \ bool" (\_\_\\_\ [51,51,51] 50) where "E\e\\T == E\In3 e \ Inr T" notation (ASCII) - wt_syntax ("_|-_::_" [51,51,51] 50) and - wt_stmt_syntax ("_|-_:<>" [51,51 ] 50) and - ty_expr_syntax ("_|-_:-_" [51,51,51] 50) and - ty_var_syntax ("_|-_:=_" [51,51,51] 50) and - ty_exprs_syntax ("_|-_:#_" [51,51,51] 50) + wt_syntax (\_|-_::_\ [51,51,51] 50) and + wt_stmt_syntax (\_|-_:<>\ [51,51 ] 50) and + ty_expr_syntax (\_|-_:-_\ [51,51,51] 50) and + ty_var_syntax (\_|-_:=_\ [51,51,51] 50) and + ty_exprs_syntax (\_|-_:#_\ [51,51,51] 50) declare not_None_eq [simp del] declare if_split [split del] if_split_asm [split del] diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Cardinals/Bounded_Set.thy --- a/src/HOL/Cardinals/Bounded_Set.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Cardinals/Bounded_Set.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ imports Cardinals begin -typedef ('a, 'k) bset ("_ set[_]" [22, 21] 21) = +typedef ('a, 'k) bset (\_ set[_]\ [22, 21] 21) = "{A :: 'a set. |A| 'a rel \ 'a rel" (infix "Osum" 60) where +definition Osum :: "'a rel \ 'a rel \ 'a rel" (infix \Osum\ 60) where "r Osum r' = r \ r' \ {(a, a'). a \ Field r \ a' \ Field r'}" -notation Osum (infix "\o" 60) +notation Osum (infix \\o\ 60) lemma Field_Osum: "Field (r \o r') = Field r \ Field r'" unfolding Osum_def Field_def by blast diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Cardinals/Ordinal_Arithmetic.thy --- a/src/HOL/Cardinals/Ordinal_Arithmetic.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ imports Wellorder_Constructions begin -definition osum :: "'a rel \ 'b rel \ ('a + 'b) rel" (infixr "+o" 70) +definition osum :: "'a rel \ 'b rel \ ('a + 'b) rel" (infixr \+o\ 70) where "r +o r' = map_prod Inl Inl ` r \ map_prod Inr Inr ` r' \ {(Inl a, Inr a') | a a' . a \ Field r \ a' \ Field r'}" @@ -161,7 +161,7 @@ lemma map_prod_ordIso: "\Well_order r; inj_on f (Field r)\ \ map_prod f f ` r =o r" by (metis dir_image_alt dir_image_ordIso ordIso_symmetric) -definition oprod :: "'a rel \ 'b rel \ ('a \ 'b) rel" (infixr "*o" 80) +definition oprod :: "'a rel \ 'b rel \ ('a \ 'b) rel" (infixr \*o\ 80) where "r *o r' = {((x1, y1), (x2, y2)). (((y1, y2) \ r' - Id \ x1 \ Field r \ x2 \ Field r) \ ((y1, y2) \ Restr Id (Field r') \ (x1, x2) \ r))}" @@ -846,7 +846,7 @@ end -notation wo_rel2.oexp (infixl "^o" 90) +notation wo_rel2.oexp (infixl \^o\ 90) lemmas oexp_def = wo_rel2.oexp_def[unfolded wo_rel2_def, OF conjI] lemmas oexp_Well_order = wo_rel2.oexp_Well_order[unfolded wo_rel2_def, OF conjI] lemmas Field_oexp = wo_rel2.Field_oexp[unfolded wo_rel2_def, OF conjI] diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Codegenerator_Test/Code_Lazy_Test.thy --- a/src/HOL/Codegenerator_Test/Code_Lazy_Test.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Codegenerator_Test/Code_Lazy_Test.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,18 +16,18 @@ value [code] "cycle ''ab''" value [code] "let x = cycle ''ab''; y = snth x 10 in x" -datatype 'a llist = LNil ("\<^bold>[\<^bold>]") | LCons (lhd: 'a) (ltl: "'a llist") (infixr "\<^bold>#" 65) +datatype 'a llist = LNil (\\<^bold>[\<^bold>]\) | LCons (lhd: 'a) (ltl: "'a llist") (infixr \\<^bold>#\ 65) subsection \Finite lazy lists\ code_lazy_type llist -no_notation lazy_llist ("_") +no_notation lazy_llist (\_\) nonterminal llist_args syntax - "" :: "'a \ llist_args" ("_") - "_llist_args" :: "'a \ llist_args \ llist_args" ("_,/ _") - "_llist" :: "llist_args => 'a list" ("\<^bold>[(_)\<^bold>]") + "" :: "'a \ llist_args" (\_\) + "_llist_args" :: "'a \ llist_args \ llist_args" (\_,/ _\) + "_llist" :: "llist_args => 'a list" (\\<^bold>[(_)\<^bold>]\) syntax_consts "_llist_args" "_llist" == lazy_llist translations @@ -77,7 +77,7 @@ subsection \Branching codatatypes\ -codatatype tree = L | Node tree tree (infix "\" 900) +codatatype tree = L | Node tree tree (infix \\\ 900) code_lazy_type tree diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Combinatorics/Perm.thy --- a/src/HOL/Combinatorics/Perm.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Combinatorics/Perm.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,7 +27,7 @@ setup_lifting type_definition_perm -notation "apply" (infixl "\$\" 999) +notation "apply" (infixl \\$\\ 999) lemma bij_apply [simp]: "bij (apply f)" @@ -731,7 +731,7 @@ subsection \Swaps\ -lift_definition swap :: "'a \ 'a \ 'a perm" ("\_ \ _\") +lift_definition swap :: "'a \ 'a \ 'a perm" (\\_ \ _\\) is "\a b. transpose a b" proof fix a b :: 'a @@ -778,7 +778,7 @@ subsection \Permutations specified by cycles\ -fun cycle :: "'a list \ 'a perm" ("\_\") +fun cycle :: "'a list \ 'a perm" (\\_\\) where "\[]\ = 1" | "\[a]\ = 1" @@ -794,16 +794,16 @@ bundle no_permutation_syntax begin - no_notation swap ("\_ \ _\") - no_notation cycle ("\_\") - no_notation "apply" (infixl "\$\" 999) + no_notation swap (\\_ \ _\\) + no_notation cycle (\\_\\) + no_notation "apply" (infixl \\$\\ 999) end bundle permutation_syntax begin - notation swap ("\_ \ _\") - notation cycle ("\_\") - notation "apply" (infixl "\$\" 999) + notation swap (\\_ \ _\\) + notation cycle (\\_\\) + notation "apply" (infixl \\$\\ 999) end unbundle no_permutation_syntax diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Complex_Analysis/Contour_Integration.thy --- a/src/HOL/Complex_Analysis/Contour_Integration.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Complex_Analysis/Contour_Integration.thy Fri Sep 20 19:51:08 2024 +0200 @@ -28,13 +28,13 @@ \ definition\<^marker>\tag important\ has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" - (infixr "has'_contour'_integral" 50) + (infixr \has'_contour'_integral\ 50) where "(f has_contour_integral i) g \ ((\x. f(g x) * vector_derivative g (at x within {0..1})) has_integral i) {0..1}" definition\<^marker>\tag important\ contour_integrable_on - (infixr "contour'_integrable'_on" 50) + (infixr \contour'_integrable'_on\ 50) where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" definition\<^marker>\tag important\ contour_integral diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Complex_Analysis/Laurent_Convergence.thy --- a/src/HOL/Complex_Analysis/Laurent_Convergence.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Complex_Analysis/Laurent_Convergence.thy Fri Sep 20 19:51:08 2024 +0200 @@ -102,7 +102,7 @@ definition\<^marker>\tag important\ has_laurent_expansion :: "(complex \ complex) \ complex fls \ bool" - (infixl "has'_laurent'_expansion" 60) + (infixl \has'_laurent'_expansion\ 60) where "(f has_laurent_expansion F) \ fls_conv_radius F > 0 \ eventually (\z. eval_fls F z = f z) (at 0)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Complex_Analysis/Meromorphic.thy --- a/src/HOL/Complex_Analysis/Meromorphic.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Complex_Analysis/Meromorphic.thy Fri Sep 20 19:51:08 2024 +0200 @@ -261,7 +261,7 @@ We will prove all of this below. \ definition%important meromorphic_on :: "(complex \ complex) \ complex set \ bool" - (infixl "(meromorphic'_on)" 50) where + (infixl \(meromorphic'_on)\ 50) where "f meromorphic_on A \ (\z\A. \F. (\w. f (z + w)) has_laurent_expansion F)" lemma meromorphic_at_iff: "f meromorphic_on {z} \ isolated_singularity_at f z \ not_essential f z" @@ -588,7 +588,7 @@ poles. We furthermore require that the function return 0 at any pole, for convenience. \ definition nicely_meromorphic_on :: "(complex \ complex) \ complex set \ bool" - (infixl "(nicely'_meromorphic'_on)" 50) + (infixl \(nicely'_meromorphic'_on)\ 50) where "f nicely_meromorphic_on A \ f meromorphic_on A \ (\z\A. (is_pole f z \ f z=0) \ f \z\ f z)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Computational_Algebra/Formal_Laurent_Series.thy --- a/src/HOL/Computational_Algebra/Formal_Laurent_Series.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Computational_Algebra/Formal_Laurent_Series.thy Fri Sep 20 19:51:08 2024 +0200 @@ -26,7 +26,7 @@ setup_lifting type_definition_fls unbundle fps_notation -notation fls_nth (infixl "$$" 75) +notation fls_nth (infixl \$$\ 75) lemmas fls_eqI = iffD1[OF fls_nth_inject, OF iffD2, OF fun_eq_iff, OF allI] @@ -4601,11 +4601,11 @@ subsection \Notation\ -no_notation fls_nth (infixl "$$" 75) +no_notation fls_nth (infixl \$$\ 75) bundle fls_notation begin -notation fls_nth (infixl "$$" 75) +notation fls_nth (infixl \$$\ 75) end end \ No newline at end of file diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Computational_Algebra/Formal_Power_Series.thy --- a/src/HOL/Computational_Algebra/Formal_Power_Series.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Computational_Algebra/Formal_Power_Series.thy Fri Sep 20 19:51:08 2024 +0200 @@ -20,7 +20,7 @@ morphisms fps_nth Abs_fps by simp -notation fps_nth (infixl "$" 75) +notation fps_nth (infixl \$\ 75) lemma expand_fps_eq: "p = q \ (\n. p $ n = q $ n)" by (simp add: fps_nth_inject [symmetric] fun_eq_iff) @@ -3633,7 +3633,7 @@ subsection \Composition\ -definition fps_compose :: "'a::semiring_1 fps \ 'a fps \ 'a fps" (infixl "oo" 55) +definition fps_compose :: "'a::semiring_1 fps \ 'a fps \ 'a fps" (infixl \oo\ 55) where "a oo b = Abs_fps (\n. sum (\i. a$i * (b^i$n)) {0..n})" lemma fps_compose_nth: "(a oo b)$n = sum (\i. a$i * (b^i$n)) {0..n}" @@ -6191,11 +6191,11 @@ qed (* TODO: Figure out better notation for this thing *) -no_notation fps_nth (infixl "$" 75) +no_notation fps_nth (infixl \$\ 75) bundle fps_notation begin -notation fps_nth (infixl "$" 75) +notation fps_nth (infixl \$\ 75) end end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Computational_Algebra/Polynomial.thy --- a/src/HOL/Computational_Algebra/Polynomial.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Computational_Algebra/Polynomial.thy Fri Sep 20 19:51:08 2024 +0200 @@ -26,7 +26,7 @@ subsection \Auxiliary: operations for lists (later) representing coefficients\ -definition cCons :: "'a::zero \ 'a list \ 'a list" (infixr "##" 65) +definition cCons :: "'a::zero \ 'a list \ 'a list" (infixr \##\ 65) where "x ## xs = (if xs = [] \ x = 0 then [] else x # xs)" lemma cCons_0_Nil_eq [simp]: "0 ## [] = []" @@ -350,9 +350,9 @@ nonterminal poly_args syntax - "" :: "'a \ poly_args" ("_") - "_poly_args" :: "'a \ poly_args \ poly_args" ("_,/ _") - "_poly" :: "poly_args \ 'a poly" ("[:(_):]") + "" :: "'a \ poly_args" (\_\) + "_poly_args" :: "'a \ poly_args \ poly_args" (\_,/ _\) + "_poly" :: "poly_args \ 'a poly" (\[:(_):]\) syntax_consts "_poly_args" "_poly" \ pCons translations @@ -2338,7 +2338,7 @@ definition pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" where "pcompose p q = fold_coeffs (\a c. [:a:] + q * c) p 0" -notation pcompose (infixl "\\<^sub>p" 71) +notation pcompose (infixl \\\<^sub>p\ 71) lemma pcompose_0 [simp]: "pcompose 0 q = 0" by (simp add: pcompose_def) @@ -6241,6 +6241,6 @@ qed (use \n > 0\ in \simp_all add: p_eq degree_power_eq\) qed -no_notation cCons (infixr "##" 65) +no_notation cCons (infixr \##\ 65) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Corec_Examples/Paper_Examples.thy --- a/src/HOL/Corec_Examples/Paper_Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Corec_Examples/Paper_Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,7 +14,7 @@ section \Examples from the introduction\ -codatatype 'a stream = SCons (shd: 'a) (stl: "'a stream") (infixr "\" 65) +codatatype 'a stream = SCons (shd: 'a) (stl: "'a stream") (infixr \\\ 65) corec "natsFrom" :: "nat \ nat stream" where "natsFrom n = n \ natsFrom (n + 1)" @@ -29,7 +29,7 @@ text \We curry the example functions in this section because infix syntax works only for curried functions.\ -corec (friend) Plus :: "nat stream \ nat stream \ nat stream" (infix "\" 67) where +corec (friend) Plus :: "nat stream \ nat stream \ nat stream" (infix \\\ 67) where "x\<^sub>1 \ x\<^sub>2 = (shd x\<^sub>1 + shd x\<^sub>2) \ (stl x\<^sub>1 \ stl x\<^sub>2)" section \Examples from section 4\ @@ -43,7 +43,7 @@ datatype 'a nelist = NEList (hd:'a) (tl:"'a list") -primrec (transfer) snoc :: "'a list \ 'a \ 'a nelist" (infix "\" 64) where +primrec (transfer) snoc :: "'a list \ 'a \ 'a nelist" (infix \\\ 64) where "[] \ a = NEList a []" |"(b # bs) \ a = NEList b (bs @ [a])" @@ -68,7 +68,7 @@ corec fibB :: "nat stream" where "fibB = (0 \ 1 \ fibB) \ (0 \ fibB)" -corec (friend) times :: "nat stream \ nat stream \ nat stream" (infix "\" 69) +corec (friend) times :: "nat stream \ nat stream \ nat stream" (infix \\\ 69) where "xs \ ys = (shd xs * shd ys) \ xs \ stl ys \ stl xs \ ys" corec (friend) exp :: "nat stream \ nat stream" @@ -106,7 +106,7 @@ corec catalan :: "nat \ nat stream" where "catalan n = (if n > 0 then catalan (n - 1) \ (0 \ catalan (n + 1)) else 1 \ catalan 1)" -corec (friend) heart :: "nat stream \ nat stream \ nat stream" (infix "\" 65) +corec (friend) heart :: "nat stream \ nat stream \ nat stream" (infix \\\ 65) where "xs \ ys = SCons (shd xs * shd ys) ((((xs \ stl ys) \ (stl xs \ ys)) \ ys) \ ys)" corec (friend) g :: "'a stream \ 'a stream" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Corec_Examples/Stream_Processor.thy --- a/src/HOL/Corec_Examples/Stream_Processor.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Corec_Examples/Stream_Processor.thy Fri Sep 20 19:51:08 2024 +0200 @@ -61,7 +61,7 @@ lemma copy_sel[simp]: "out copy = Get (\a. Put a copy)" by (subst copy.code; simp) -corecursive sp_comp (infixl "oo" 65) where +corecursive sp_comp (infixl \oo\ 65) where "sp oo sp' = (case (out sp, out sp') of (Put b sp, _) \ In (Put b (sp oo sp')) | (Get f, Put b sp) \ In (f b) oo sp diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Data_Structures/Selection.thy --- a/src/HOL/Data_Structures/Selection.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Data_Structures/Selection.thy Fri Sep 20 19:51:08 2024 +0200 @@ -397,7 +397,7 @@ begin lemma size_median_of_medians_aux: - fixes R :: "'a :: linorder \ 'a \ bool" (infix "\" 50) + fixes R :: "'a :: linorder \ 'a \ bool" (infix \\\ 50) assumes R: "R \ {(<), (>)}" shows "size {#y \# mset xs. y \ M#} \ nat \0.7 * length xs + 3\" proof - @@ -406,8 +406,8 @@ text \We then split that multiset into those groups whose medians is less than @{term M} and the rest.\ - define Y_small ("Y\<^sub>\") where "Y\<^sub>\ = filter_mset (\ys. median ys \ M) (mset (chop 5 xs))" - define Y_big ("Y\<^sub>\") where "Y\<^sub>\ = filter_mset (\ys. \(median ys \ M)) (mset (chop 5 xs))" + define Y_small (\Y\<^sub>\\) where "Y\<^sub>\ = filter_mset (\ys. median ys \ M) (mset (chop 5 xs))" + define Y_big (\Y\<^sub>\\) where "Y\<^sub>\ = filter_mset (\ys. \(median ys \ M)) (mset (chop 5 xs))" have "m = size (mset (chop 5 xs))" by (simp add: m_def) also have "mset (chop 5 xs) = Y\<^sub>\ + Y\<^sub>\" unfolding Y_small_def Y_big_def by (rule multiset_partition) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Data_Structures/Tree23.thy --- a/src/HOL/Data_Structures/Tree23.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Data_Structures/Tree23.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,9 +10,9 @@ fixes height :: "'a \ nat" datatype 'a tree23 = - Leaf ("\\") | - Node2 "'a tree23" 'a "'a tree23" ("\_, _, _\") | - Node3 "'a tree23" 'a "'a tree23" 'a "'a tree23" ("\_, _, _, _, _\") + Leaf (\\\\) | + Node2 "'a tree23" 'a "'a tree23" (\\_, _, _\\) | + Node3 "'a tree23" 'a "'a tree23" 'a "'a tree23" (\\_, _, _, _, _\\) fun inorder :: "'a tree23 \ 'a list" where "inorder Leaf = []" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Data_Structures/Tree234.thy --- a/src/HOL/Data_Structures/Tree234.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Data_Structures/Tree234.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,11 +10,11 @@ fixes height :: "'a \ nat" datatype 'a tree234 = - Leaf ("\\") | - Node2 "'a tree234" 'a "'a tree234" ("\_, _, _\") | - Node3 "'a tree234" 'a "'a tree234" 'a "'a tree234" ("\_, _, _, _, _\") | + Leaf (\\\\) | + Node2 "'a tree234" 'a "'a tree234" (\\_, _, _\\) | + Node3 "'a tree234" 'a "'a tree234" 'a "'a tree234" (\\_, _, _, _, _\\) | Node4 "'a tree234" 'a "'a tree234" 'a "'a tree234" 'a "'a tree234" - ("\_, _, _, _, _, _, _\") + (\\_, _, _, _, _, _, _\\) fun inorder :: "'a tree234 \ 'a list" where "inorder Leaf = []" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Datatype_Examples/Derivation_Trees/Parallel_Composition.thy --- a/src/HOL/Datatype_Examples/Derivation_Trees/Parallel_Composition.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Datatype_Examples/Derivation_Trees/Parallel_Composition.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,9 +11,9 @@ imports DTree begin -no_notation plus_class.plus (infixl "+" 65) +no_notation plus_class.plus (infixl \+\ 65) -consts Nplus :: "N \ N \ N" (infixl "+" 60) +consts Nplus :: "N \ N \ N" (infixl \+\ 60) axiomatization where Nplus_comm: "(a::N) + b = b + (a::N)" @@ -32,7 +32,7 @@ definition par :: "dtree \ dtree \ dtree" where "par \ unfold par_r par_c" -abbreviation par_abbr (infixr "\" 80) where "tr1 \ tr2 \ par (tr1, tr2)" +abbreviation par_abbr (infixr \\\ 80) where "tr1 \ tr2 \ par (tr1, tr2)" lemma finite_par_c: "finite (par_c (tr1, tr2))" unfolding par_c.simps apply(rule finite_UnI) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Datatype_Examples/Derivation_Trees/Prelim.thy --- a/src/HOL/Datatype_Examples/Derivation_Trees/Prelim.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Datatype_Examples/Derivation_Trees/Prelim.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,14 +11,14 @@ imports "HOL-Library.FSet" begin -notation BNF_Def.convol ("\(_,/ _)\") +notation BNF_Def.convol (\\(_,/ _)\\) declare fset_to_fset[simp] lemma fst_snd_convol_o[simp]: "\fst o s, snd o s\ = s" apply(rule ext) by (simp add: convol_def) -abbreviation sm_abbrev (infix "\" 60) +abbreviation sm_abbrev (infix \\\ 60) where "f \ g \ Sum_Type.map_sum f g" lemma map_sum_InlD: "(f \ g) z = Inl x \ \y. z = Inl y \ f y = x" @@ -27,7 +27,7 @@ lemma map_sum_InrD: "(f \ g) z = Inr x \ \y. z = Inr y \ g y = x" by (cases z) auto -abbreviation case_sum_abbrev ("[[_,_]]" 800) +abbreviation case_sum_abbrev (\[[_,_]]\ 800) where "[[f,g]] \ Sum_Type.case_sum f g" lemma Inl_oplus_elim: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Datatype_Examples/Koenig.thy --- a/src/HOL/Datatype_Examples/Koenig.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Datatype_Examples/Koenig.thy Fri Sep 20 19:51:08 2024 +0200 @@ -90,11 +90,11 @@ (* some more stream theorems *) -primcorec plus :: "nat stream \ nat stream \ nat stream" (infixr "\" 66) where +primcorec plus :: "nat stream \ nat stream \ nat stream" (infixr \\\ 66) where "shd (plus xs ys) = shd xs + shd ys" | "stl (plus xs ys) = plus (stl xs) (stl ys)" -definition scalar :: "nat \ nat stream \ nat stream" (infixr "\" 68) where +definition scalar :: "nat \ nat stream \ nat stream" (infixr \\\ 68) where [simp]: "scalar n = smap (\x. n * x)" primcorec ones :: "nat stream" where "ones = 1 ## ones" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Datatype_Examples/Milner_Tofte.thy --- a/src/HOL/Datatype_Examples/Milner_Tofte.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Datatype_Examples/Milner_Tofte.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,16 +27,16 @@ datatype Ex = e_const (e_const_fst: Const) | e_var ExVar -| e_fn ExVar Ex ("fn _ => _" [0,51] 1000) -| e_fix ExVar ExVar Ex ("fix _ ( _ ) = _" [0,51,51] 1000) -| e_app Ex Ex ("_ @@ _" [51,51] 1000) +| e_fn ExVar Ex (\fn _ => _\ [0,51] 1000) +| e_fix ExVar ExVar Ex (\fix _ ( _ ) = _\ [0,51,51] 1000) +| e_app Ex Ex (\_ @@ _\ [51,51] 1000) datatype Ty = t_const TyConst -| t_fun Ty Ty ("_ -> _" [51,51] 1000) +| t_fun Ty Ty (\_ -> _\ [51,51] 1000) datatype 'a genClos = - clos_mk ExVar Ex "ExVar \ 'a" ("\_ , _ , _\" [0,0,0] 1000) + clos_mk ExVar Ex "ExVar \ 'a" (\\_ , _ , _\\ [0,0,0] 1000) codatatype Val = v_const Const @@ -48,7 +48,7 @@ axiomatization c_app :: "[Const, Const] => Const" and - isof :: "[Const, Ty] => bool" ("_ isof _" [36,36] 50) where + isof :: "[Const, Ty] => bool" (\_ isof _\ [36,36] 50) where isof_app: "\c1 isof t1 -> t2; c2 isof t1\ \ c_app c1 c2 isof t2" text \The dynamic semantics is defined inductively by a set of inference @@ -58,7 +58,7 @@ as the least fixpoint of the functor eval_fun below. From this definition introduction rules and a strong elimination (induction) rule can be derived.\ -inductive eval :: "[ValEnv, Ex, Val] => bool" ("_ \ _ ---> _" [36,0,36] 50) where +inductive eval :: "[ValEnv, Ex, Val] => bool" (\_ \ _ ---> _\ [36,0,36] 50) where eval_const: "ve \ e_const c ---> v_const c" | eval_var2: "ev \ dom ve \ ve \ e_var ev ---> the (ve ev)" | eval_fn: "ve \ fn ev => e ---> v_clos \ev, e, ve\" @@ -74,7 +74,7 @@ semantics. The relation te |- e ===> t express the expression e has the type t in the type environment te.\ -inductive elab :: "[TyEnv, Ex, Ty] => bool" ("_ \ _ ===> _" [36,0,36] 50) where +inductive elab :: "[TyEnv, Ex, Ty] => bool" (\_ \ _ ===> _\ [36,0,36] 50) where elab_const: "c isof ty \ te \ e_const c ===> ty" | elab_var: "x \ dom te \ te \ e_var x ===> the (te x)" | elab_fn: "te(x \ ty1) \ e ===> ty2 \ te \ fn x => e ===> ty1 -> ty2" @@ -91,11 +91,11 @@ (* The original correspondence relation *) -abbreviation isof_env :: "[ValEnv,TyEnv] => bool" ("_ isofenv _") where +abbreviation isof_env :: "[ValEnv,TyEnv] => bool" (\_ isofenv _\) where "ve isofenv te \ (dom(ve) = dom(te) \ (\x. x \ dom ve \ (\c. the (ve x) = v_const(c) \ c isof the (te x))))" -coinductive hasty :: "[Val, Ty] => bool" ("_ hasty _" [36,36] 50) where +coinductive hasty :: "[Val, Ty] => bool" (\_ hasty _\ [36,36] 50) where hasty_const: "c isof t \ v_const c hasty t" | hasty_clos: "\te \ fn ev => e ===> t; dom(ve) = dom(te) \ (\x. x \ dom ve --> the (ve x) hasty the (te x)); cl = \ev,e,ve\\ \ v_clos cl hasty t" @@ -105,7 +105,7 @@ "v_const c hasty t" "v_clos \xm , em , evm\ hasty u -> t" -definition hasty_env :: "[ValEnv, TyEnv] => bool" ("_ hastyenv _ " [36,36] 35) where +definition hasty_env :: "[ValEnv, TyEnv] => bool" (\_ hastyenv _ \ [36,36] 35) where [simp]: "ve hastyenv te \ (dom(ve) = dom(te) \ (\x. x \ dom ve --> the (ve x) hasty the (te x)))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Datatype_Examples/Regex_ACI.thy --- a/src/HOL/Datatype_Examples/Regex_ACI.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Datatype_Examples/Regex_ACI.thy Fri Sep 20 19:51:08 2024 +0200 @@ -4,7 +4,7 @@ datatype 'a rexp = Zero | Eps | Atom 'a | Alt "'a rexp" "'a rexp" | Conc "'a rexp" "'a rexp" | Star "'a rexp" -inductive ACI (infix "~" 64) where +inductive ACI (infix \~\ 64) where a1: "Alt (Alt r s) t ~ Alt r (Alt s t)" | a2: "Alt r (Alt s t) ~ Alt (Alt r s) t" | c: "Alt r s ~ Alt s r" @@ -16,7 +16,7 @@ declare ACI.intros[intro] -abbreviation ACIcl (infix "~~" 64) where "(~~) \ equivclp (~)" +abbreviation ACIcl (infix \~~\ 64) where "(~~) \ equivclp (~)" lemma eq_set_preserves_inter: fixes eq set @@ -305,7 +305,7 @@ intro: equivpI reflpI sympI transpI strong_confluentp_imp_confluentp[OF strong_confluentp_ACI]) -inductive ACIEQ (infix "\" 64) where +inductive ACIEQ (infix \\\ 64) where a: "Alt (Alt r s) t \ Alt r (Alt s t)" | c: "Alt r s \ Alt s r" | i: "Alt r r \ r" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Datatype_Examples/Regex_ACIDZ.thy --- a/src/HOL/Datatype_Examples/Regex_ACIDZ.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Datatype_Examples/Regex_ACIDZ.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,7 +17,7 @@ | "distribute t (Conc r s) = Conc (distribute s r) t" | "distribute t r = Conc r t" -inductive ACIDZ (infix "~" 64) where +inductive ACIDZ (infix \~\ 64) where a1: "Alt (Alt r s) t ~ Alt r (Alt s t)" | a2: "Alt r (Alt s t) ~ Alt (Alt r s) t" | c: "Alt r s ~ Alt s r" @@ -34,7 +34,7 @@ declare ACIDZ.intros[intro] -abbreviation ACIDZcl (infix "~~" 64) where "(~~) \ equivclp (~)" +abbreviation ACIDZcl (infix \~~\ 64) where "(~~) \ equivclp (~)" lemma map_rexp_distribute[simp]: "map_rexp f (distribute t r) = distribute (map_rexp f t) (map_rexp f r)" by (induct r arbitrary: t) auto @@ -399,7 +399,7 @@ by unfold_locales (auto intro: ACIDZcl_map_respects rexp.map_cong simp: rexp.map_id rexp.set_map finite_set_rexp) -inductive ACIDZEQ (infix "\" 64) where +inductive ACIDZEQ (infix \\\ 64) where a: "Alt (Alt r s) t \ Alt r (Alt s t)" | c: "Alt r s \ Alt s r" | i: "Alt r r \ r" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Datatype_Examples/Stream_Processor.thy --- a/src/HOL/Datatype_Examples/Stream_Processor.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Datatype_Examples/Stream_Processor.thy Fri Sep 20 19:51:08 2024 +0200 @@ -61,7 +61,7 @@ function sp\<^sub>\_comp :: "('a, 'b, 'c) sp\<^sub>\ \ ('d, 'a, ('d, 'a) sp\<^sub>\) sp\<^sub>\ \ ('d, 'b, 'c \ ('d, 'a) sp\<^sub>\) sp\<^sub>\" - (infixl "o\<^sub>\" 65) + (infixl \o\<^sub>\\ 65) where "Put b sp o\<^sub>\ fsp = Put b (sp, In fsp)" | "Get f o\<^sub>\ Put b sp = f b o\<^sub>\ out sp" @@ -70,7 +70,7 @@ termination by (relation "lex_prod sub sub") auto -primcorec sp\<^sub>\_comp (infixl "o\<^sub>\" 65) where +primcorec sp\<^sub>\_comp (infixl \o\<^sub>\\ 65) where "out (sp o\<^sub>\ sp') = map_sp\<^sub>\ id (\(sp, sp'). sp o\<^sub>\ sp') (out sp o\<^sub>\ out sp')" lemma run\<^sub>\_sp\<^sub>\_comp: "run\<^sub>\ (sp o\<^sub>\ sp') = run\<^sub>\ sp o run\<^sub>\ sp'" @@ -99,11 +99,11 @@ "sp\<^sub>\_comp2R f (Put b sp) = f b (out sp)" | "sp\<^sub>\_comp2R f (Get h) = Get (sp\<^sub>\_comp2R f o h)" -primrec sp\<^sub>\_comp2 (infixl "o\<^sup>*\<^sub>\" 65) where +primrec sp\<^sub>\_comp2 (infixl \o\<^sup>*\<^sub>\\ 65) where "Put b sp o\<^sup>*\<^sub>\ fsp = Put b (sp, In fsp)" | "Get f o\<^sup>*\<^sub>\ fsp = sp\<^sub>\_comp2R ((o\<^sup>*\<^sub>\) o f) fsp" -primcorec sp\<^sub>\_comp2 (infixl "o\<^sup>*\<^sub>\" 65) where +primcorec sp\<^sub>\_comp2 (infixl \o\<^sup>*\<^sub>\\ 65) where "out (sp o\<^sup>*\<^sub>\ sp') = map_sp\<^sub>\ id (\(sp, sp'). sp o\<^sup>*\<^sub>\ sp') (out sp o\<^sup>*\<^sub>\ out sp')" lemma run\<^sub>\_sp\<^sub>\_comp2: "run\<^sub>\ (sp o\<^sup>*\<^sub>\ sp') = run\<^sub>\ sp o run\<^sub>\ sp'" @@ -175,7 +175,7 @@ bnf_axiomatization ('a, 'b) F for map: F -notation BNF_Def.convol ("\(_,/ _)\") +notation BNF_Def.convol (\\(_,/ _)\\) definition \ :: "('p,'a) F * 'b \ ('p,'a * 'b) F" where "\ xb = F id \id, \ a. (snd xb)\ (fst xb)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Decision_Procs/Algebra_Aux.thy --- a/src/HOL/Decision_Procs/Algebra_Aux.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Decision_Procs/Algebra_Aux.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,10 +8,10 @@ imports "HOL-Algebra.Ring" begin -definition of_natural :: "('a, 'm) ring_scheme \ nat \ 'a" ("\_\\<^sub>\\") +definition of_natural :: "('a, 'm) ring_scheme \ nat \ 'a" (\\_\\<^sub>\\\) where "\n\\<^sub>\\<^bsub>R\<^esub> = ((\\<^bsub>R\<^esub>) \\<^bsub>R\<^esub> ^^ n) \\<^bsub>R\<^esub>" -definition of_integer :: "('a, 'm) ring_scheme \ int \ 'a" ("\_\\") +definition of_integer :: "('a, 'm) ring_scheme \ int \ 'a" (\\_\\\) where "\i\\<^bsub>R\<^esub> = (if 0 \ i then \nat i\\<^sub>\\<^bsub>R\<^esub> else \\<^bsub>R\<^esub> \nat (- i)\\<^sub>\\<^bsub>R\<^esub>)" context ring @@ -313,7 +313,7 @@ by (auto simp add: l_minus r_minus) qed -definition m_div :: "('a, 'b) ring_scheme \ 'a \ 'a \ 'a" (infixl "\\" 70) +definition m_div :: "('a, 'b) ring_scheme \ 'a \ 'a \ 'a" (infixl \\\\ 70) where "x \\<^bsub>G\<^esub> y = (if y = \\<^bsub>G\<^esub> then \\<^bsub>G\<^esub> else x \\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y)" context field diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Decision_Procs/Approximation.thy --- a/src/HOL/Decision_Procs/Approximation.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Decision_Procs/Approximation.thy Fri Sep 20 19:51:08 2024 +0200 @@ -1145,47 +1145,47 @@ bundle floatarith_notation begin -notation Add ("Add") -notation Minus ("Minus") -notation Mult ("Mult") -notation Inverse ("Inverse") -notation Cos ("Cos") -notation Arctan ("Arctan") -notation Abs ("Abs") -notation Max ("Max") -notation Min ("Min") -notation Pi ("Pi") -notation Sqrt ("Sqrt") -notation Exp ("Exp") -notation Powr ("Powr") -notation Ln ("Ln") -notation Power ("Power") -notation Floor ("Floor") -notation Var ("Var") -notation Num ("Num") +notation Add (\Add\) +notation Minus (\Minus\) +notation Mult (\Mult\) +notation Inverse (\Inverse\) +notation Cos (\Cos\) +notation Arctan (\Arctan\) +notation Abs (\Abs\) +notation Max (\Max\) +notation Min (\Min\) +notation Pi (\Pi\) +notation Sqrt (\Sqrt\) +notation Exp (\Exp\) +notation Powr (\Powr\) +notation Ln (\Ln\) +notation Power (\Power\) +notation Floor (\Floor\) +notation Var (\Var\) +notation Num (\Num\) end bundle no_floatarith_notation begin -no_notation Add ("Add") -no_notation Minus ("Minus") -no_notation Mult ("Mult") -no_notation Inverse ("Inverse") -no_notation Cos ("Cos") -no_notation Arctan ("Arctan") -no_notation Abs ("Abs") -no_notation Max ("Max") -no_notation Min ("Min") -no_notation Pi ("Pi") -no_notation Sqrt ("Sqrt") -no_notation Exp ("Exp") -no_notation Powr ("Powr") -no_notation Ln ("Ln") -no_notation Power ("Power") -no_notation Floor ("Floor") -no_notation Var ("Var") -no_notation Num ("Num") +no_notation Add (\Add\) +no_notation Minus (\Minus\) +no_notation Mult (\Mult\) +no_notation Inverse (\Inverse\) +no_notation Cos (\Cos\) +no_notation Arctan (\Arctan\) +no_notation Abs (\Abs\) +no_notation Max (\Max\) +no_notation Min (\Min\) +no_notation Pi (\Pi\) +no_notation Sqrt (\Sqrt\) +no_notation Exp (\Exp\) +no_notation Powr (\Powr\) +no_notation Ln (\Ln\) +no_notation Power (\Power\) +no_notation Floor (\Floor\) +no_notation Var (\Var\) +no_notation Num (\Num\) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Decision_Procs/Commutative_Ring.thy --- a/src/HOL/Decision_Procs/Commutative_Ring.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Decision_Procs/Commutative_Ring.thy Fri Sep 20 19:51:08 2024 +0200 @@ -106,7 +106,7 @@ text \Defining the basic ring operations on normalized polynomials\ -function add :: "pol \ pol \ pol" (infixl "\+\" 65) +function add :: "pol \ pol \ pol" (infixl \\+\\ 65) where "Pc a \+\ Pc b = Pc (a + b)" | "Pc c \+\ Pinj i P = Pinj i (P \+\ Pc c)" @@ -132,7 +132,7 @@ by pat_completeness auto termination by (relation "measure (\(x, y). size x + size y)") auto -function mul :: "pol \ pol \ pol" (infixl "\*\" 70) +function mul :: "pol \ pol \ pol" (infixl \\*\\ 70) where "Pc a \*\ Pc b = Pc (a * b)" | "Pc c \*\ Pinj i P = @@ -174,7 +174,7 @@ | "neg (PX P x Q) = PX (neg P) x (neg Q)" text \Subtraction\ -definition sub :: "pol \ pol \ pol" (infixl "\-\" 65) +definition sub :: "pol \ pol \ pol" (infixl \\-\\ 65) where "sub P Q = P \+\ neg Q" text \Square for Fast Exponentiation\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Decision_Procs/Dense_Linear_Order.thy --- a/src/HOL/Decision_Procs/Dense_Linear_Order.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Decision_Procs/Dense_Linear_Order.thy Fri Sep 20 19:51:08 2024 +0200 @@ -438,10 +438,10 @@ begin notation - less_eq ("'(\')") and - less_eq ("(_/ \ _)" [51, 51] 50) and - less ("'(\')") and - less ("(_/ \ _)" [51, 51] 50) + less_eq (\'(\')\) and + less_eq (\(_/ \ _)\ [51, 51] 50) and + less (\'(\')\) and + less (\(_/ \ _)\ [51, 51] 50) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Decision_Procs/MIR.thy Fri Sep 20 19:51:08 2024 +0200 @@ -31,7 +31,7 @@ lemmas dvd_period = zdvd_period (* The Divisibility relation between reals *) -definition rdvd:: "real \ real \ bool" (infixl "rdvd" 50) +definition rdvd:: "real \ real \ bool" (infixl \rdvd\ 50) where "x rdvd y \ (\k::int. y = x * real_of_int k)" lemma int_rdvd_real: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy --- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,8 +13,8 @@ datatype poly = C Num | Bound nat | Add poly poly | Sub poly poly | Mul poly poly| Neg poly| Pw poly nat| CN poly nat poly -abbreviation poly_0 :: "poly" ("0\<^sub>p") where "0\<^sub>p \ C (0\<^sub>N)" -abbreviation poly_p :: "int \ poly" ("'((_)')\<^sub>p") where "(i)\<^sub>p \ C (i)\<^sub>N" +abbreviation poly_0 :: "poly" (\0\<^sub>p\) where "0\<^sub>p \ C (0\<^sub>N)" +abbreviation poly_p :: "int \ poly" (\'((_)')\<^sub>p\) where "(i)\<^sub>p \ C (i)\<^sub>N" subsection\Boundedness, substitution and all that\ @@ -116,7 +116,7 @@ declare if_cong[fundef_cong del] declare let_cong[fundef_cong del] -fun polyadd :: "poly \ poly \ poly" (infixl "+\<^sub>p" 60) +fun polyadd :: "poly \ poly \ poly" (infixl \+\<^sub>p\ 60) where "polyadd (C c) (C c') = C (c +\<^sub>N c')" | "polyadd (C c) (CN c' n' p') = CN (polyadd (C c) c') n' p'" @@ -131,16 +131,16 @@ in if pp' = 0\<^sub>p then cc' else CN cc' n pp')" | "polyadd a b = Add a b" -fun polyneg :: "poly \ poly" ("~\<^sub>p") +fun polyneg :: "poly \ poly" (\~\<^sub>p\) where "polyneg (C c) = C (~\<^sub>N c)" | "polyneg (CN c n p) = CN (polyneg c) n (polyneg p)" | "polyneg a = Neg a" -definition polysub :: "poly \ poly \ poly" (infixl "-\<^sub>p" 60) +definition polysub :: "poly \ poly \ poly" (infixl \-\<^sub>p\ 60) where "p -\<^sub>p q = polyadd p (polyneg q)" -fun polymul :: "poly \ poly \ poly" (infixl "*\<^sub>p" 60) +fun polymul :: "poly \ poly \ poly" (infixl \*\<^sub>p\ 60) where "polymul (C c) (C c') = C (c *\<^sub>N c')" | "polymul (C c) (CN c' n' p') = @@ -166,7 +166,7 @@ d = polymul q q in if even n then d else polymul p d)" -abbreviation poly_pow :: "poly \ nat \ poly" (infixl "^\<^sub>p" 60) +abbreviation poly_pow :: "poly \ nat \ poly" (infixl \^\<^sub>p\ 60) where "a ^\<^sub>p k \ polypow k a" function polynate :: "poly \ poly" @@ -245,7 +245,7 @@ | "Ipoly bs (Pw t n) = Ipoly bs t ^ n" | "Ipoly bs (CN c n p) = Ipoly bs c + (bs!n) * Ipoly bs p" -abbreviation Ipoly_syntax :: "poly \ 'a list \'a::{field_char_0,power}" ("\_\\<^sub>p\<^bsup>_\<^esup>") +abbreviation Ipoly_syntax :: "poly \ 'a list \'a::{field_char_0,power}" (\\_\\<^sub>p\<^bsup>_\<^esup>\) where "\p\\<^sub>p\<^bsup>bs\<^esup> \ Ipoly bs p" lemma Ipoly_CInt: "Ipoly bs (C (i, 1)) = of_int i" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Examples/Adhoc_Overloading_Examples.thy --- a/src/HOL/Examples/Adhoc_Overloading_Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Examples/Adhoc_Overloading_Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -176,7 +176,7 @@ this end we introduce a constant \PERMUTE\ together with convenient infix syntax.\ -consts PERMUTE :: "'a perm \ 'b \ 'b" (infixr "\" 75) +consts PERMUTE :: "'a perm \ 'b \ 'b" (infixr \\\ 75) text \Then we add a locale for types \<^typ>\'b\ that support appliciation of permutations.\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Examples/Coherent.thy --- a/src/HOL/Examples/Coherent.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Examples/Coherent.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,8 +12,8 @@ subsection \Equivalence of two versions of Pappus' Axiom\ no_notation - comp (infixl "o" 55) and - relcomp (infixr "O" 75) + comp (infixl \o\ 55) and + relcomp (infixr \O\ 75) lemma p1p2: assumes "col a b c l \ col d e f m" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Cfun.thy --- a/src/HOL/HOLCF/Cfun.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Cfun.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,17 +16,17 @@ definition "cfun = {f::'a \ 'b. cont f}" -cpodef ('a, 'b) cfun ("(_ \/ _)" [1, 0] 0) = "cfun :: ('a \ 'b) set" +cpodef ('a, 'b) cfun (\(_ \/ _)\ [1, 0] 0) = "cfun :: ('a \ 'b) set" by (auto simp: cfun_def intro: cont_const adm_cont) type_notation (ASCII) - cfun (infixr "->" 0) + cfun (infixr \->\ 0) notation (ASCII) - Rep_cfun ("(_$/_)" [999,1000] 999) + Rep_cfun (\(_$/_)\ [999,1000] 999) notation - Rep_cfun ("(_\/_)" [999,1000] 999) + Rep_cfun (\(_\/_)\ [999,1000] 999) subsection \Syntax for continuous lambda abstraction\ @@ -47,10 +47,10 @@ text \Syntax for nested abstractions\ syntax (ASCII) - "_Lambda" :: "[cargs, logic] \ logic" ("(3LAM _./ _)" [1000, 10] 10) + "_Lambda" :: "[cargs, logic] \ logic" (\(3LAM _./ _)\ [1000, 10] 10) syntax - "_Lambda" :: "[cargs, logic] \ logic" ("(3\ _./ _)" [1000, 10] 10) + "_Lambda" :: "[cargs, logic] \ logic" (\(3\ _./ _)\ [1000, 10] 10) syntax_consts "_Lambda" \ Abs_cfun @@ -417,7 +417,7 @@ definition cfcomp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" where oo_def: "cfcomp = (\ f g x. f\(g\x))" -abbreviation cfcomp_syn :: "['b \ 'c, 'a \ 'b] \ 'a \ 'c" (infixr "oo" 100) +abbreviation cfcomp_syn :: "['b \ 'c, 'a \ 'b] \ 'a \ 'c" (infixr \oo\ 100) where "f oo g == cfcomp\f\g" lemma ID1 [simp]: "ID\x = x" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Completion.thy --- a/src/HOL/HOLCF/Completion.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Completion.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ subsection \Ideals over a preorder\ locale preorder = - fixes r :: "'a::type \ 'a \ bool" (infix "\" 50) + fixes r :: "'a::type \ 'a \ bool" (infix \\\ 50) assumes r_refl: "x \ x" assumes r_trans: "\x \ y; y \ z\ \ x \ z" begin diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/ConvexPD.thy --- a/src/HOL/HOLCF/ConvexPD.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/ConvexPD.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ subsection \Basis preorder\ definition - convex_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix "\\" 50) where + convex_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix \\\\ 50) where "convex_le = (\u v. u \\ v \ u \\ v)" lemma convex_le_refl [simp]: "t \\ t" @@ -119,7 +119,7 @@ subsection \Type definition\ -typedef 'a convex_pd ("('(_')\)") = +typedef 'a convex_pd (\('(_')\)\) = "{S::'a pd_basis set. convex_le.ideal S}" by (rule convex_le.ex_ideal) @@ -175,14 +175,14 @@ abbreviation convex_add :: "'a convex_pd \ 'a convex_pd \ 'a convex_pd" - (infixl "\\" 65) where + (infixl \\\\ 65) where "xs \\ ys == convex_plus\xs\ys" nonterminal convex_pd_args syntax - "" :: "logic \ convex_pd_args" ("_") - "_convex_pd_args" :: "logic \ convex_pd_args \ convex_pd_args" ("_,/ _") - "_convex_pd" :: "convex_pd_args \ logic" ("{_}\") + "" :: "logic \ convex_pd_args" (\_\) + "_convex_pd_args" :: "logic \ convex_pd_args \ convex_pd_args" (\_,/ _\) + "_convex_pd" :: "convex_pd_args \ logic" (\{_}\\) syntax_consts "_convex_pd_args" "_convex_pd" == convex_add translations @@ -350,7 +350,7 @@ syntax "_convex_bind" :: "[logic, logic, logic] \ logic" - ("(3\\_\_./ _)" [0, 0, 10] 10) + (\(3\\_\_./ _)\ [0, 0, 10] 10) syntax_consts "_convex_bind" == convex_bind diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/FOCUS/Buffer.thy --- a/src/HOL/HOLCF/FOCUS/Buffer.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/FOCUS/Buffer.thy Fri Sep 20 19:51:08 2024 +0200 @@ -26,10 +26,10 @@ typedecl D datatype - M = Md D | Mreq ("\") + M = Md D | Mreq (\\\) datatype - State = Sd D | Snil ("\") + State = Sd D | Snil (\\\) type_synonym SPF11 = "M fstream \ D fstream" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/FOCUS/Fstream.thy --- a/src/HOL/HOLCF/FOCUS/Fstream.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/FOCUS/Fstream.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,20 +25,20 @@ "fsfilter A = (sfilter\(flift2 (\x. x\A)))" abbreviation - emptystream :: "'a fstream" ("<>") where + emptystream :: "'a fstream" (\<>\) where "<> == \" abbreviation - fscons' :: "'a \ 'a fstream \ 'a fstream" ("(_\_)" [66,65] 65) where + fscons' :: "'a \ 'a fstream \ 'a fstream" (\(_\_)\ [66,65] 65) where "a\s == fscons a\s" abbreviation - fsfilter' :: "'a set \ 'a fstream \ 'a fstream" ("(_\_)" [64,63] 63) where + fsfilter' :: "'a set \ 'a fstream \ 'a fstream" (\(_\_)\ [64,63] 63) where "A\s == fsfilter A\s" notation (ASCII) - fscons' ("(_~>_)" [66,65] 65) and - fsfilter' ("(_'(C')_)" [64,63] 63) + fscons' (\(_~>_)\ [66,65] 65) and + fsfilter' (\(_'(C')_)\ [64,63] 63) lemma Def_maximal: "a = Def d \ a\b \ b = Def d" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/FOCUS/Fstreams.thy --- a/src/HOL/HOLCF/FOCUS/Fstreams.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/FOCUS/Fstreams.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,7 +15,7 @@ type_synonym 'a fstream = "('a lift) stream" definition - fsingleton :: "'a => 'a fstream" ("<_>" [1000] 999) where + fsingleton :: "'a => 'a fstream" (\<_>\ [1000] 999) where fsingleton_def2: "fsingleton = (%a. Def a && UU)" definition @@ -40,15 +40,15 @@ abbreviation - emptystream :: "'a fstream" ("<>") where + emptystream :: "'a fstream" (\<>\) where "<> == \" abbreviation - fsfilter' :: "'a set \ 'a fstream \ 'a fstream" ("(_\_)" [64,63] 63) where + fsfilter' :: "'a set \ 'a fstream \ 'a fstream" (\(_\_)\ [64,63] 63) where "A\s == fsfilter A\s" notation (ASCII) - fsfilter' ("(_'(C')_)" [64,63] 63) + fsfilter' (\(_'(C')_)\ [64,63] 63) lemma ft_fsingleton[simp]: "ft\() = Def a" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Fix.thy --- a/src/HOL/HOLCF/Fix.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Fix.thy Fri Sep 20 19:51:08 2024 +0200 @@ -48,11 +48,11 @@ text \Binder syntax for \<^term>\fix\\ -abbreviation fix_syn :: "('a \ 'a) \ 'a" (binder "\ " 10) +abbreviation fix_syn :: "('a \ 'a) \ 'a" (binder \\ \ 10) where "fix_syn (\x. f x) \ fix\(\ x. f x)" notation (ASCII) - fix_syn (binder "FIX " 10) + fix_syn (binder \FIX \ 10) text \Properties of \<^term>\fix\\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Fixrec.thy --- a/src/HOL/HOLCF/Fixrec.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Fixrec.thy Fri Sep 20 19:51:08 2024 +0200 @@ -73,7 +73,7 @@ "mplus = (\ m1 m2. sscase\(\ _. m2)\(\ _. m1)\(Rep_match m1))" abbreviation - mplus_syn :: "['a match, 'a match] \ 'a match" (infixr "+++" 65) where + mplus_syn :: "['a match, 'a match] \ 'a match" (infixr \+++\ 65) where "m1 +++ m2 == mplus\m1\m2" text \rewrite rules for mplus\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IMP/HoareEx.thy --- a/src/HOL/HOLCF/IMP/HoareEx.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IMP/HoareEx.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ type_synonym assn = "state \ bool" definition - hoare_valid :: "[assn, com, assn] \ bool" ("|= {(1_)}/ (_)/ {(1_)}" 50) where + hoare_valid :: "[assn, com, assn] \ bool" (\|= {(1_)}/ (_)/ {(1_)}\ 50) where "|= {P} c {Q} = (\s t. P s \ D c\(Discr s) = Def t \ Q t)" lemma WHILE_rule_sound: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IOA/Automata.thy --- a/src/HOL/HOLCF/IOA/Automata.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IOA/Automata.thy Fri Sep 20 19:51:08 2024 +0200 @@ -26,7 +26,7 @@ definition trans_of :: "('a, 's) ioa \ ('a, 's) transition set" where "trans_of = fst \ snd \ snd" -abbreviation trans_of_syn ("_ \_\_\ _" [81, 81, 81, 81] 100) +abbreviation trans_of_syn (\_ \_\_\ _\ [81, 81, 81, 81] 100) where "s \a\A\ t \ (s, a, t) \ trans_of A" definition wfair_of :: "('a, 's) ioa \ 'a set set" @@ -91,7 +91,7 @@ (outputs a1 \ outputs a2), (internals a1 \ internals a2)))" -definition par :: "('a, 's) ioa \ ('a, 't) ioa \ ('a, 's * 't) ioa" (infixr "\" 10) +definition par :: "('a, 's) ioa \ ('a, 't) ioa \ ('a, 's * 't) ioa" (infixr \\\ 10) where "(A \ B) = (asig_comp (asig_of A) (asig_of B), {pr. fst pr \ starts_of A \ snd pr \ starts_of B}, diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IOA/NTP/Multiset.thy --- a/src/HOL/HOLCF/IOA/NTP/Multiset.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IOA/NTP/Multiset.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ consts - emptym :: "'a multiset" ("{|}") + emptym :: "'a multiset" (\{|}\) addm :: "['a multiset, 'a] => 'a multiset" delm :: "['a multiset, 'a] => 'a multiset" countm :: "['a multiset, 'a => bool] => nat" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IOA/Pred.thy --- a/src/HOL/HOLCF/IOA/Pred.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IOA/Pred.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,22 +12,22 @@ type_synonym 'a predicate = "'a \ bool" -definition satisfies :: "'a \ 'a predicate \ bool" ("_ \ _" [100, 9] 8) +definition satisfies :: "'a \ 'a predicate \ bool" (\_ \ _\ [100, 9] 8) where "(s \ P) \ P s" -definition valid :: "'a predicate \ bool" ("\ _" [9] 8) +definition valid :: "'a predicate \ bool" (\\ _\ [9] 8) where "(\ P) \ (\s. (s \ P))" -definition Not :: "'a predicate \ 'a predicate" ("\<^bold>\ _" [40] 40) +definition Not :: "'a predicate \ 'a predicate" (\\<^bold>\ _\ [40] 40) where NOT_def: "Not P s \ \ P s" -definition AND :: "'a predicate \ 'a predicate \ 'a predicate" (infixr "\<^bold>\" 35) +definition AND :: "'a predicate \ 'a predicate \ 'a predicate" (infixr \\<^bold>\\ 35) where "(P \<^bold>\ Q) s \ P s \ Q s" -definition OR :: "'a predicate \ 'a predicate \ 'a predicate" (infixr "\<^bold>\" 30) +definition OR :: "'a predicate \ 'a predicate \ 'a predicate" (infixr \\<^bold>\\ 30) where "(P \<^bold>\ Q) s \ P s \ Q s" -definition IMPLIES :: "'a predicate \ 'a predicate \ 'a predicate" (infixr "\<^bold>\" 25) +definition IMPLIES :: "'a predicate \ 'a predicate \ 'a predicate" (infixr \\<^bold>\\ 25) where "(P \<^bold>\ Q) s \ P s \ Q s" end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IOA/Seq.thy --- a/src/HOL/HOLCF/IOA/Seq.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IOA/Seq.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,7 +10,7 @@ default_sort pcpo -domain (unsafe) 'a seq = nil ("nil") | cons (HD :: 'a) (lazy TL :: "'a seq") (infixr "##" 65) +domain (unsafe) 'a seq = nil (\nil\) | cons (HD :: 'a) (lazy TL :: "'a seq") (infixr \##\ 65) inductive Finite :: "'a seq \ bool" where @@ -108,7 +108,7 @@ sconc_nil: "sconc \ nil \ y = y" | sconc_cons': "x \ UU \ sconc \ (x ## xs) \ y = x ## (sconc \ xs \ y)" -abbreviation sconc_syn :: "'a seq \ 'a seq \ 'a seq" (infixr "@@" 65) +abbreviation sconc_syn :: "'a seq \ 'a seq \ 'a seq" (infixr \@@\ 65) where "xs @@ ys \ sconc \ xs \ ys" lemma sconc_UU [simp]: "UU @@ y = UU" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IOA/Sequence.thy --- a/src/HOL/HOLCF/IOA/Sequence.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IOA/Sequence.thy Fri Sep 20 19:51:08 2024 +0200 @@ -63,7 +63,7 @@ UU \ UU | Def y \ (if P y then x ## (h \ xs) else h \ xs))))" -abbreviation Consq_syn ("(_/\_)" [66, 65] 65) +abbreviation Consq_syn (\(_/\_)\ [66, 65] 65) where "a \ s \ Consq a \ s" @@ -71,10 +71,10 @@ nonterminal llist_args syntax - "" :: "'a \ llist_args" ("_") - "_list_args" :: "'a \ llist_args \ llist_args" ("_,/ _") - "_totlist" :: "llist_args \ 'a Seq" ("[(_)!]") - "_partlist" :: "llist_args \ 'a Seq" ("[(_)?]") + "" :: "'a \ llist_args" (\_\) + "_list_args" :: "'a \ llist_args \ llist_args" (\_,/ _\) + "_totlist" :: "llist_args \ 'a Seq" (\[(_)!]\) + "_partlist" :: "llist_args \ 'a Seq" (\[(_)?]\) syntax_consts "_totlist" "_partlist" \ Consq translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IOA/TL.thy --- a/src/HOL/HOLCF/IOA/TL.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IOA/TL.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,17 +12,17 @@ type_synonym 'a temporal = "'a Seq predicate" -definition validT :: "'a Seq predicate \ bool" ("\<^bold>\ _" [9] 8) +definition validT :: "'a Seq predicate \ bool" (\\<^bold>\ _\ [9] 8) where "(\<^bold>\ P) \ (\s. s \ UU \ s \ nil \ (s \ P))" definition unlift :: "'a lift \ 'a" where "unlift x = (case x of Def y \ y)" -definition Init :: "'a predicate \ 'a temporal" ("\_\" [0] 1000) +definition Init :: "'a predicate \ 'a temporal" (\\_\\ [0] 1000) where "Init P s = P (unlift (HD \ s))" \ \this means that for \nil\ and \UU\ the effect is unpredictable\ -definition Next :: "'a temporal \ 'a temporal" ("\(_)" [80] 80) +definition Next :: "'a temporal \ 'a temporal" (\\(_)\ [80] 80) where "(\P) s \ (if TL \ s = UU \ TL \ s = nil then P s else P (TL \ s))" definition suffix :: "'a Seq \ 'a Seq \ bool" @@ -31,13 +31,13 @@ definition tsuffix :: "'a Seq \ 'a Seq \ bool" where "tsuffix s2 s \ s2 \ nil \ s2 \ UU \ suffix s2 s" -definition Box :: "'a temporal \ 'a temporal" ("\(_)" [80] 80) +definition Box :: "'a temporal \ 'a temporal" (\\(_)\ [80] 80) where "(\P) s \ (\s2. tsuffix s2 s \ P s2)" -definition Diamond :: "'a temporal \ 'a temporal" ("\(_)" [80] 80) +definition Diamond :: "'a temporal \ 'a temporal" (\\(_)\ [80] 80) where "\P = (\<^bold>\ (\(\<^bold>\ P)))" -definition Leadsto :: "'a temporal \ 'a temporal \ 'a temporal" (infixr "\" 22) +definition Leadsto :: "'a temporal \ 'a temporal \ 'a temporal" (infixr \\\ 22) where "(P \ Q) = (\(P \<^bold>\ (\Q)))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IOA/TLS.thy --- a/src/HOL/HOLCF/IOA/TLS.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IOA/TLS.thy Fri Sep 20 19:51:08 2024 +0200 @@ -42,7 +42,7 @@ definition ex2seq :: "('a, 's) execution \ ('a option, 's) transition Seq" where "ex2seq ex = (ex2seqC \ (mkfin (snd ex))) (fst ex)" -definition temp_sat :: "('a, 's) execution \ ('a, 's) ioa_temp \ bool" (infixr "\" 22) +definition temp_sat :: "('a, 's) execution \ ('a, 's) ioa_temp \ bool" (infixr \\\ 22) where "(ex \ P) \ ((ex2seq ex) \ P)" definition validTE :: "('a, 's) ioa_temp \ bool" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/IOA/Traces.thy --- a/src/HOL/HOLCF/IOA/Traces.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/IOA/Traces.thy Fri Sep 20 19:51:08 2024 +0200 @@ -129,7 +129,7 @@ subsubsection \Notions of implementation\ -definition ioa_implements :: "('a, 's1) ioa \ ('a, 's2) ioa \ bool" (infixr "=<|" 12) +definition ioa_implements :: "('a, 's1) ioa \ ('a, 's2) ioa \ bool" (infixr \=<|\ 12) where "(ioa1 =<| ioa2) \ (inputs (asig_of ioa1) = inputs (asig_of ioa2) \ outputs (asig_of ioa1) = outputs (asig_of ioa2)) \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Library/Stream.thy --- a/src/HOL/HOLCF/Library/Stream.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Library/Stream.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,7 +10,7 @@ default_sort pcpo -domain (unsafe) 'a stream = scons (ft::'a) (lazy rt::"'a stream") (infixr "&&" 65) +domain (unsafe) 'a stream = scons (ft::'a) (lazy rt::"'a stream") (infixr \&&\ 65) definition smap :: "('a \ 'b) \ 'a stream \ 'b stream" where @@ -22,7 +22,7 @@ If p\x then x && h\p\xs else h\p\xs)" definition - slen :: "'a stream \ enat" ("#_" [1000] 1000) where + slen :: "'a stream \ enat" (\#_\ [1000] 1000) where "#s = (if stream_finite s then enat (LEAST n. stream_take n\s = s) else \)" @@ -37,7 +37,7 @@ "i_th = (\i s. ft\(i_rt i s))" definition - sconc :: "'a stream \ 'a stream \ 'a stream" (infixr "ooo" 65) where + sconc :: "'a stream \ 'a stream \ 'a stream" (infixr \ooo\ 65) where "s1 ooo s2 = (case #s1 of enat n \ (SOME s. (stream_take n\s = s1) \ (i_rt n s = s2)) | \ \ s1)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Lift.thy --- a/src/HOL/HOLCF/Lift.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Lift.thy Fri Sep 20 19:51:08 2024 +0200 @@ -85,7 +85,7 @@ subsection \Further operations\ definition - flift1 :: "('a \ 'b::pcpo) \ ('a lift \ 'b)" (binder "FLIFT " 10) where + flift1 :: "('a \ 'b::pcpo) \ ('a lift \ 'b)" (binder \FLIFT \ 10) where "flift1 = (\f. (\ x. case_lift \ f x))" translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/LowerPD.thy --- a/src/HOL/HOLCF/LowerPD.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/LowerPD.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ subsection \Basis preorder\ definition - lower_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix "\\" 50) where + lower_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix \\\\ 50) where "lower_le = (\u v. \x\Rep_pd_basis u. \y\Rep_pd_basis v. x \ y)" lemma lower_le_refl [simp]: "t \\ t" @@ -74,7 +74,7 @@ subsection \Type definition\ -typedef 'a lower_pd ("('(_')\)") = +typedef 'a lower_pd (\('(_')\)\) = "{S::'a pd_basis set. lower_le.ideal S}" by (rule lower_le.ex_ideal) @@ -130,14 +130,14 @@ abbreviation lower_add :: "'a lower_pd \ 'a lower_pd \ 'a lower_pd" - (infixl "\\" 65) where + (infixl \\\\ 65) where "xs \\ ys == lower_plus\xs\ys" nonterminal lower_pd_args syntax - "" :: "logic \ lower_pd_args" ("_") - "_lower_pd_args" :: "logic \ lower_pd_args \ lower_pd_args" ("_,/ _") - "_lower_pd" :: "lower_pd_args \ logic" ("{_}\") + "" :: "logic \ lower_pd_args" (\_\) + "_lower_pd_args" :: "logic \ lower_pd_args \ lower_pd_args" (\_,/ _\) + "_lower_pd" :: "lower_pd_args \ logic" (\{_}\\) syntax_consts "_lower_pd_args" "_lower_pd" == lower_add translations @@ -344,7 +344,7 @@ syntax "_lower_bind" :: "[logic, logic, logic] \ logic" - ("(3\\_\_./ _)" [0, 0, 10] 10) + (\(3\\_\_./ _)\ [0, 0, 10] 10) syntax_consts "_lower_bind" == lower_bind diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Pcpo.thy --- a/src/HOL/HOLCF/Pcpo.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Pcpo.thy Fri Sep 20 19:51:08 2024 +0200 @@ -136,7 +136,7 @@ assumes least: "\x. \y. x \ y" begin -definition bottom :: "'a" ("\") +definition bottom :: "'a" (\\\) where "bottom = (THE x. \y. x \ y)" lemma minimal [iff]: "\ \ x" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Porder.thy --- a/src/HOL/HOLCF/Porder.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Porder.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,16 +18,16 @@ begin notation (ASCII) - below (infix "<<" 50) + below (infix \<<\ 50) notation - below (infix "\" 50) + below (infix \\\ 50) -abbreviation not_below :: "'a \ 'a \ bool" (infix "\" 50) +abbreviation not_below :: "'a \ 'a \ bool" (infix \\\ 50) where "not_below x y \ \ below x y" notation (ASCII) - not_below (infix "~<<" 50) + not_below (infix \~<<\ 50) lemma below_eq_trans: "a \ b \ b = c \ a \ c" by (rule subst) @@ -71,7 +71,7 @@ subsection \Upper bounds\ -definition is_ub :: "'a set \ 'a \ bool" (infix "<|" 55) +definition is_ub :: "'a set \ 'a \ bool" (infix \<|\ 55) where "S <| x \ (\y\S. y \ x)" lemma is_ubI: "(\x. x \ S \ x \ u) \ S <| u" @@ -104,7 +104,7 @@ subsection \Least upper bounds\ -definition is_lub :: "'a set \ 'a \ bool" (infix "<<|" 55) +definition is_lub :: "'a set \ 'a \ bool" (infix \<<|\ 55) where "S <<| x \ S <| x \ (\u. S <| u \ x \ u)" definition lub :: "'a set \ 'a" @@ -113,10 +113,10 @@ end syntax (ASCII) - "_BLub" :: "[pttrn, 'a set, 'b] \ 'b" ("(3LUB _:_./ _)" [0,0, 10] 10) + "_BLub" :: "[pttrn, 'a set, 'b] \ 'b" (\(3LUB _:_./ _)\ [0,0, 10] 10) syntax - "_BLub" :: "[pttrn, 'a set, 'b] \ 'b" ("(3\_\_./ _)" [0,0, 10] 10) + "_BLub" :: "[pttrn, 'a set, 'b] \ 'b" (\(3\_\_./ _)\ [0,0, 10] 10) syntax_consts "_BLub" \ lub @@ -127,11 +127,11 @@ context po begin -abbreviation Lub (binder "\" 10) +abbreviation Lub (binder \\\ 10) where "\n. t n \ lub (range t)" notation (ASCII) - Lub (binder "LUB " 10) + Lub (binder \LUB \ 10) text \access to some definition as inference rule\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Representable.thy --- a/src/HOL/HOLCF/Representable.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Representable.thy Fri Sep 20 19:51:08 2024 +0200 @@ -31,7 +31,7 @@ assumes predomain_ep: "ep_pair liftemb liftprj" assumes cast_liftdefl: "cast\(liftdefl TYPE('a)) = liftemb oo liftprj" -syntax "_LIFTDEFL" :: "type \ logic" ("(1LIFTDEFL/(1'(_')))") +syntax "_LIFTDEFL" :: "type \ logic" (\(1LIFTDEFL/(1'(_')))\) syntax_consts "_LIFTDEFL" \ liftdefl translations "LIFTDEFL('t)" \ "CONST liftdefl TYPE('t)" @@ -51,7 +51,7 @@ assumes liftprj_eq: "liftprj = u_map\prj" assumes liftdefl_eq: "liftdefl TYPE('a) = liftdefl_of\(defl TYPE('a))" -syntax "_DEFL" :: "type \ logic" ("(1DEFL/(1'(_')))") +syntax "_DEFL" :: "type \ logic" (\(1DEFL/(1'(_')))\) syntax_consts "_DEFL" \ defl translations "DEFL('t)" \ "CONST defl TYPE('t)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Sfun.thy --- a/src/HOL/HOLCF/Sfun.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Sfun.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,11 +8,11 @@ imports Cfun begin -pcpodef ('a, 'b) sfun (infixr "\!" 0) = "{f :: 'a \ 'b. f\\ = \}" +pcpodef ('a, 'b) sfun (infixr \\!\ 0) = "{f :: 'a \ 'b. f\\ = \}" by simp_all type_notation (ASCII) - sfun (infixr "->!" 0) + sfun (infixr \->!\ 0) text \TODO: Define nice syntax for abstraction, application.\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Sprod.thy --- a/src/HOL/HOLCF/Sprod.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Sprod.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,14 +16,14 @@ definition "sprod = {p::'a \ 'b. p = \ \ (fst p \ \ \ snd p \ \)}" -pcpodef ('a, 'b) sprod ("(_ \/ _)" [21,20] 20) = "sprod :: ('a \ 'b) set" +pcpodef ('a, 'b) sprod (\(_ \/ _)\ [21,20] 20) = "sprod :: ('a \ 'b) set" by (simp_all add: sprod_def) instance sprod :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin by (rule typedef_chfin [OF type_definition_sprod below_sprod_def]) type_notation (ASCII) - sprod (infixr "**" 20) + sprod (infixr \**\ 20) subsection \Definitions of constants\ @@ -42,9 +42,9 @@ nonterminal stuple_args syntax - "" :: "logic \ stuple_args" ("_") - "_stuple_args" :: "logic \ stuple_args \ stuple_args" ("_,/ _") - "_stuple" :: "[logic, stuple_args] \ logic" ("(1'(:_,/ _:'))") + "" :: "logic \ stuple_args" (\_\) + "_stuple_args" :: "logic \ stuple_args \ stuple_args" (\_,/ _\) + "_stuple" :: "[logic, stuple_args] \ logic" (\(1'(:_,/ _:'))\) syntax_consts "_stuple_args" "_stuple" \ spair translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Ssum.thy --- a/src/HOL/HOLCF/Ssum.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Ssum.thy Fri Sep 20 19:51:08 2024 +0200 @@ -19,14 +19,14 @@ (fst p = TT \ fst (snd p) \ \ \ snd (snd p) = \) \ (fst p = FF \ fst (snd p) = \ \ snd (snd p) \ \)}" -pcpodef ('a, 'b) ssum ("(_ \/ _)" [21, 20] 20) = "ssum :: (tr \ 'a \ 'b) set" +pcpodef ('a, 'b) ssum (\(_ \/ _)\ [21, 20] 20) = "ssum :: (tr \ 'a \ 'b) set" by (simp_all add: ssum_def) instance ssum :: ("{chfin,pcpo}", "{chfin,pcpo}") chfin by (rule typedef_chfin [OF type_definition_ssum below_ssum_def]) type_notation (ASCII) - ssum (infixr "++" 10) + ssum (infixr \++\ 10) subsection \Definitions of constructors\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Tr.thy --- a/src/HOL/HOLCF/Tr.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Tr.thy Fri Sep 20 19:51:08 2024 +0200 @@ -63,7 +63,7 @@ definition tr_case :: "'a \ 'a \ tr \ 'a" where "tr_case = (\ t e (Def b). if b then t else e)" -abbreviation cifte_syn :: "[tr, 'c, 'c] \ 'c" ("(If (_)/ then (_)/ else (_))" [0, 0, 60] 60) +abbreviation cifte_syn :: "[tr, 'c, 'c] \ 'c" (\(If (_)/ then (_)/ else (_))\ [0, 0, 60] 60) where "If b then e1 else e2 \ tr_case\e1\e2\b" translations @@ -82,13 +82,13 @@ definition trand :: "tr \ tr \ tr" where andalso_def: "trand = (\ x y. If x then y else FF)" -abbreviation andalso_syn :: "tr \ tr \ tr" ("_ andalso _" [36,35] 35) +abbreviation andalso_syn :: "tr \ tr \ tr" (\_ andalso _\ [36,35] 35) where "x andalso y \ trand\x\y" definition tror :: "tr \ tr \ tr" where orelse_def: "tror = (\ x y. If x then TT else y)" -abbreviation orelse_syn :: "tr \ tr \ tr" ("_ orelse _" [31,30] 30) +abbreviation orelse_syn :: "tr \ tr \ tr" (\_ orelse _\ [31,30] 30) where "x orelse y \ tror\x\y" definition neg :: "tr \ tr" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Tutorial/Domain_ex.thy --- a/src/HOL/HOLCF/Tutorial/Domain_ex.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Tutorial/Domain_ex.thy Fri Sep 20 19:51:08 2024 +0200 @@ -34,14 +34,14 @@ text \Mixfix declarations can be given for data constructors.\ -domain d5 = d5a | d5b (lazy "d5") "d5" (infixl ":#:" 70) +domain d5 = d5a | d5b (lazy "d5") "d5" (infixl \:#:\ 70) lemma "d5a \ x :#: y :#: z" by simp text \Mixfix declarations can also be given for type constructors.\ -domain ('a, 'b) lazypair (infixl ":*:" 25) = - lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl ":*:" 75) +domain ('a, 'b) lazypair (infixl \:*:\ 25) = + lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl \:*:\ 75) lemma "\p::('a :*: 'b). p \ lfst\p :*: lsnd\p" by (rule allI, case_tac p, simp_all) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Universal.thy --- a/src/HOL/HOLCF/Universal.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Universal.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,7 +8,7 @@ imports Bifinite Completion "HOL-Library.Nat_Bijection" begin -no_notation binomial (infix "choose" 64) +no_notation binomial (infix \choose\ 64) subsection \Basis for universal domain\ @@ -973,6 +973,6 @@ hide_const (open) node -notation binomial (infixl "choose" 65) +notation binomial (infixl \choose\ 65) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/Up.thy --- a/src/HOL/HOLCF/Up.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/Up.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,7 +14,7 @@ subsection \Definition of new type for lifting\ -datatype 'a u ("(_\<^sub>\)" [1000] 999) = Ibottom | Iup 'a +datatype 'a u (\(_\<^sub>\)\ [1000] 999) = Ibottom | Iup 'a primrec Ifup :: "('a \ 'b::pcpo) \ 'a u \ 'b" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/UpperPD.thy --- a/src/HOL/HOLCF/UpperPD.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/UpperPD.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ subsection \Basis preorder\ definition - upper_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix "\\" 50) where + upper_le :: "'a pd_basis \ 'a pd_basis \ bool" (infix \\\\ 50) where "upper_le = (\u v. \y\Rep_pd_basis v. \x\Rep_pd_basis u. x \ y)" lemma upper_le_refl [simp]: "t \\ t" @@ -72,7 +72,7 @@ subsection \Type definition\ -typedef 'a upper_pd ("('(_')\)") = +typedef 'a upper_pd (\('(_')\)\) = "{S::'a pd_basis set. upper_le.ideal S}" by (rule upper_le.ex_ideal) @@ -128,14 +128,14 @@ abbreviation upper_add :: "'a upper_pd \ 'a upper_pd \ 'a upper_pd" - (infixl "\\" 65) where + (infixl \\\\ 65) where "xs \\ ys == upper_plus\xs\ys" nonterminal upper_pd_args syntax - "" :: "logic \ upper_pd_args" ("_") - "_upper_pd_args" :: "logic \ upper_pd_args \ upper_pd_args" ("_,/ _") - "_upper_pd" :: "upper_pd_args \ logic" ("{_}\") + "" :: "logic \ upper_pd_args" (\_\) + "_upper_pd_args" :: "logic \ upper_pd_args \ upper_pd_args" (\_,/ _\) + "_upper_pd" :: "upper_pd_args \ logic" (\{_}\\) syntax_consts "_upper_pd_args" "_upper_pd" == upper_add translations @@ -337,7 +337,7 @@ syntax "_upper_bind" :: "[logic, logic, logic] \ logic" - ("(3\\_\_./ _)" [0, 0, 10] 10) + (\(3\\_\_./ _)\ [0, 0, 10] 10) syntax_consts "_upper_bind" == upper_bind diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/ex/Concurrency_Monad.thy --- a/src/HOL/HOLCF/ex/Concurrency_Monad.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/ex/Concurrency_Monad.thy Fri Sep 20 19:51:08 2024 +0200 @@ -200,7 +200,7 @@ lemma zipR_strict2 [simp]: "zipR\f\r\\ = \" by (fixrec_simp, cases r, simp_all) -abbreviation apR (infixl "\" 70) +abbreviation apR (infixl \\\ 70) where "a \ b \ zipR\ID\a\b" text \Proofs that \zipR\ satisfies the applicative functor laws:\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/ex/Letrec.thy --- a/src/HOL/HOLCF/ex/Letrec.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/ex/Letrec.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,12 +15,12 @@ nonterminal recbinds and recbindt and recbind syntax - "_recbind" :: "[logic, logic] \ recbind" ("(2_ =/ _)" 10) - "" :: "recbind \ recbindt" ("_") - "_recbindt" :: "[recbind, recbindt] \ recbindt" ("_,/ _") - "" :: "recbindt \ recbinds" ("_") - "_recbinds" :: "[recbindt, recbinds] \ recbinds" ("_;/ _") - "_Letrec" :: "[recbinds, logic] \ logic" ("(Letrec (_)/ in (_))" 10) + "_recbind" :: "[logic, logic] \ recbind" (\(2_ =/ _)\ 10) + "" :: "recbind \ recbindt" (\_\) + "_recbindt" :: "[recbind, recbindt] \ recbindt" (\_,/ _\) + "" :: "recbindt \ recbinds" (\_\) + "_recbinds" :: "[recbindt, recbinds] \ recbinds" (\_;/ _\) + "_Letrec" :: "[recbinds, logic] \ logic" (\(Letrec (_)/ in (_))\ 10) syntax_consts "_recbind" "_recbinds" "_recbindt" "_Letrec" == CLetrec diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/HOLCF/ex/Pattern_Match.thy --- a/src/HOL/HOLCF/ex/Pattern_Match.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/HOLCF/ex/Pattern_Match.thy Fri Sep 20 19:51:08 2024 +0200 @@ -28,7 +28,7 @@ "fatbar = (\ a b x. a\x +++ b\x)" abbreviation - fatbar_syn :: "['a \ 'b match, 'a \ 'b match] \ 'a \ 'b match" (infixr "\" 60) where + fatbar_syn :: "['a \ 'b match, 'a \ 'b match] \ 'a \ 'b match" (infixr \\\ 60) where "m1 \ m2 == fatbar\m1\m2" lemma fatbar1: "m\x = \ \ (m \ ms)\x = \" @@ -103,14 +103,14 @@ nonterminal Case_pat and Case_syn and Cases_syn syntax - "_Case_syntax":: "['a, Cases_syn] => 'b" ("(Case _ of/ _)" 10) - "_Case1" :: "[Case_pat, 'b] => Case_syn" ("(2_ \/ _)" 10) - "" :: "Case_syn => Cases_syn" ("_") - "_Case2" :: "[Case_syn, Cases_syn] => Cases_syn" ("_/ | _") - "_strip_positions" :: "'a => Case_pat" ("_") + "_Case_syntax":: "['a, Cases_syn] => 'b" (\(Case _ of/ _)\ 10) + "_Case1" :: "[Case_pat, 'b] => Case_syn" (\(2_ \/ _)\ 10) + "" :: "Case_syn => Cases_syn" (\_\) + "_Case2" :: "[Case_syn, Cases_syn] => Cases_syn" (\_/ | _\) + "_strip_positions" :: "'a => Case_pat" (\_\) syntax (ASCII) - "_Case1" :: "[Case_pat, 'b] => Case_syn" ("(2_ =>/ _)" 10) + "_Case1" :: "[Case_pat, 'b] => Case_syn" (\(2_ =>/ _)\ 10) translations "_Case_syntax x ms" == "CONST cases\(ms\x)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hahn_Banach/Bounds.thy --- a/src/HOL/Hahn_Banach/Bounds.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hahn_Banach/Bounds.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,7 +15,7 @@ lemmas [elim?] = lub.least lub.upper -definition the_lub :: "'a::order set \ 'a" ("\_" [90] 90) +definition the_lub :: "'a::order set \ 'a" (\\_\ [90] 90) where "the_lub A = The (lub A)" lemma the_lub_equality [elim?]: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hahn_Banach/Function_Norm.thy --- a/src/HOL/Hahn_Banach/Function_Norm.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hahn_Banach/Function_Norm.thy Fri Sep 20 19:51:08 2024 +0200 @@ -21,13 +21,13 @@ \ locale continuous = linearform + - fixes norm :: "_ \ real" ("\_\") + fixes norm :: "_ \ real" (\\_\\) assumes bounded: "\c. \x \ V. \f x\ \ c * \x\" declare continuous.intro [intro?] continuous_axioms.intro [intro?] lemma continuousI [intro]: - fixes norm :: "_ \ real" ("\_\") + fixes norm :: "_ \ real" (\\_\\) assumes "linearform V f" assumes r: "\x. x \ V \ \f x\ \ c * \x\" shows "continuous V f norm" @@ -69,9 +69,9 @@ \ locale fn_norm = - fixes norm :: "_ \ real" ("\_\") + fixes norm :: "_ \ real" (\\_\\) fixes B defines "B V f \ {0} \ {\f x\ / \x\ | x. x \ 0 \ x \ V}" - fixes fn_norm ("\_\\_" [0, 1000] 999) + fixes fn_norm (\\_\\_\ [0, 1000] 999) defines "\f\\V \ \(B V f)" locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hahn_Banach/Hahn_Banach.thy --- a/src/HOL/Hahn_Banach/Hahn_Banach.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hahn_Banach/Hahn_Banach.thy Fri Sep 20 19:51:08 2024 +0200 @@ -338,9 +338,9 @@ \ theorem norm_Hahn_Banach: - fixes V and norm ("\_\") + fixes V and norm (\\_\\) fixes B defines "\V f. B V f \ {0} \ {\f x\ / \x\ | x. x \ 0 \ x \ V}" - fixes fn_norm ("\_\\_" [0, 1000] 999) + fixes fn_norm (\\_\\_\ [0, 1000] 999) defines "\V f. \f\\V \ \(B V f)" assumes E_norm: "normed_vectorspace E norm" and FE: "subspace F E" and linearform: "linearform F f" and "continuous F f norm" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hahn_Banach/Normed_Space.thy --- a/src/HOL/Hahn_Banach/Normed_Space.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hahn_Banach/Normed_Space.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,7 +18,7 @@ locale seminorm = fixes V :: "'a::{minus, plus, zero, uminus} set" - fixes norm :: "'a \ real" ("\_\") + fixes norm :: "'a \ real" (\\_\\) assumes ge_zero [iff?]: "x \ V \ 0 \ \x\" and abs_homogenous [iff?]: "x \ V \ \a \ x\ = \a\ * \x\" and subadditive [iff?]: "x \ V \ y \ V \ \x + y\ \ \x\ + \y\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hahn_Banach/Subspace.thy --- a/src/HOL/Hahn_Banach/Subspace.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hahn_Banach/Subspace.thy Fri Sep 20 19:51:08 2024 +0200 @@ -23,7 +23,7 @@ and mult_closed [iff]: "x \ U \ a \ x \ U" notation (symbols) - subspace (infix "\" 50) + subspace (infix \\\ 50) declare vectorspace.intro [intro?] subspace.intro [intro?] diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hahn_Banach/Vector_Space.thy --- a/src/HOL/Hahn_Banach/Vector_Space.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hahn_Banach/Vector_Space.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,7 +17,7 @@ \ consts - prod :: "real \ 'a::{plus,minus,zero} \ 'a" (infixr "\" 70) + prod :: "real \ 'a::{plus,minus,zero} \ 'a" (infixr \\\ 70) subsection \Vector space laws\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare/HeapSyntax.thy --- a/src/HOL/Hoare/HeapSyntax.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare/HeapSyntax.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,11 +13,11 @@ syntax "_refupdate" :: "('a \ 'b) \ 'a ref \ 'b \ ('a \ 'b)" - ("_/'((_ \ _)')" [1000,0] 900) + (\_/'((_ \ _)')\ [1000,0] 900) "_fassign" :: "'a ref => id => 'v => 's com" - ("(2_^._ :=/ _)" [70,1000,65] 61) + (\(2_^._ :=/ _)\ [70,1000,65] 61) "_faccess" :: "'a ref => ('a ref \ 'v) => 'v" - ("_^._" [65,1000] 65) + (\_^._\ [65,1000] 65) translations "f(r \ v)" == "f(CONST addr r := v)" "p^.f := e" => "f := f(p \ e)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare/HeapSyntaxAbort.thy --- a/src/HOL/Hoare/HeapSyntaxAbort.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare/HeapSyntaxAbort.thy Fri Sep 20 19:51:08 2024 +0200 @@ -21,11 +21,11 @@ syntax "_refupdate" :: "('a \ 'b) \ 'a ref \ 'b \ ('a \ 'b)" - ("_/'((_ \ _)')" [1000,0] 900) + (\_/'((_ \ _)')\ [1000,0] 900) "_fassign" :: "'a ref => id => 'v => 's com" - ("(2_^._ :=/ _)" [70,1000,65] 61) + (\(2_^._ :=/ _)\ [70,1000,65] 61) "_faccess" :: "'a ref => ('a ref \ 'v) => 'v" - ("_^._" [65,1000] 65) + (\_^._\ [65,1000] 65) translations "_refupdate f r v" == "f(CONST addr r := v)" "p^.f := e" => "(p \ CONST Null) \ (f := _refupdate f p e)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare/Hoare_Logic.thy --- a/src/HOL/Hoare/Hoare_Logic.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare/Hoare_Logic.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,7 +27,7 @@ | Cond "'a bexp" "'a com" "'a com" | While "'a bexp" "'a com" -abbreviation annskip ("SKIP") where "SKIP == Basic id" +abbreviation annskip (\SKIP\) where "SKIP == Basic id" type_synonym 'a sem = "'a => 'a => bool" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare/Hoare_Logic_Abort.thy --- a/src/HOL/Hoare/Hoare_Logic_Abort.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy Fri Sep 20 19:51:08 2024 +0200 @@ -21,7 +21,7 @@ | Cond "'a bexp" "'a com" "'a com" | While "'a bexp" "'a com" -abbreviation annskip ("SKIP") where "SKIP == Basic id" +abbreviation annskip (\SKIP\) where "SKIP == Basic id" type_synonym 'a sem = "'a option => 'a option => bool" @@ -187,8 +187,8 @@ \ \Special syntax for guarded statements and guarded array updates:\ syntax - "_guarded_com" :: "bool \ 'a com \ 'a com" ("(2_ \/ _)" 71) - "_array_update" :: "'a list \ nat \ 'a \ 'a com" ("(2_[_] :=/ _)" [70, 65] 61) + "_guarded_com" :: "bool \ 'a com \ 'a com" (\(2_ \/ _)\ 71) + "_array_update" :: "'a list \ nat \ 'a \ 'a com" (\(2_[_] :=/ _)\ [70, 65] 61) translations "P \ c" \ "IF P THEN c ELSE CONST Abort FI" "a[i] := v" \ "(i < CONST length a) \ (a := CONST list_update a i v)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare/Hoare_Syntax.thy --- a/src/HOL/Hoare/Hoare_Syntax.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare/Hoare_Syntax.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,10 +10,10 @@ begin syntax - "_assign" :: "idt \ 'b \ 'com" ("(2_ :=/ _)" [70, 65] 61) - "_Seq" :: "'com \ 'com \ 'com" ("(_;/ _)" [61,60] 60) - "_Cond" :: "'bexp \ 'com \ 'com \ 'com" ("(IF _/ THEN _ / ELSE _/ FI)" [0,0,0] 61) - "_While" :: "'bexp \ 'assn \ 'var \ 'com \ 'com" ("(WHILE _/ INV {(_)} / VAR {(_)} //DO _ /OD)" [0,0,0,0] 61) + "_assign" :: "idt \ 'b \ 'com" (\(2_ :=/ _)\ [70, 65] 61) + "_Seq" :: "'com \ 'com \ 'com" (\(_;/ _)\ [61,60] 60) + "_Cond" :: "'bexp \ 'com \ 'com \ 'com" (\(IF _/ THEN _ / ELSE _/ FI)\ [0,0,0] 61) + "_While" :: "'bexp \ 'assn \ 'var \ 'com \ 'com" (\(WHILE _/ INV {(_)} / VAR {(_)} //DO _ /OD)\ [0,0,0,0] 61) text \The \VAR {_}\ syntax supports two variants: \<^item> \VAR {x = t}\ where \t::nat\ is the decreasing expression, @@ -23,17 +23,17 @@ \ syntax - "_While0" :: "'bexp \ 'assn \ 'com \ 'com" ("(1WHILE _/ INV {_} //DO _ /OD)" [0,0,0] 61) + "_While0" :: "'bexp \ 'assn \ 'com \ 'com" (\(1WHILE _/ INV {_} //DO _ /OD)\ [0,0,0] 61) text \The \_While0\ syntax is translated into the \_While\ syntax with the trivial variant 0. This is ok because partial correctness proofs do not make use of the variant.\ syntax - "_hoare_vars" :: "[idts, 'assn,'com, 'assn] \ bool" ("VARS _// {_} // _ // {_}" [0,0,55,0] 50) - "_hoare_vars_tc" :: "[idts, 'assn, 'com, 'assn] \ bool" ("VARS _// [_] // _ // [_]" [0,0,55,0] 50) + "_hoare_vars" :: "[idts, 'assn,'com, 'assn] \ bool" (\VARS _// {_} // _ // {_}\ [0,0,55,0] 50) + "_hoare_vars_tc" :: "[idts, 'assn, 'com, 'assn] \ bool" (\VARS _// [_] // _ // [_]\ [0,0,55,0] 50) syntax (output) - "_hoare" :: "['assn, 'com, 'assn] \ bool" ("({_}//_//{_})" [0,55,0] 50) - "_hoare_tc" :: "['assn, 'com, 'assn] \ bool" ("([_]//_//[_])" [0,55,0] 50) + "_hoare" :: "['assn, 'com, 'assn] \ bool" (\({_}//_//{_})\ [0,55,0] 50) + "_hoare_tc" :: "['assn, 'com, 'assn] \ bool" (\([_]//_//[_])\ [0,55,0] 50) text \Completeness requires(?) the ability to refer to an outer variant in an inner invariant. Thus we need to abstract over a variable equated with the variant, the \x\ in \VAR {x = t}\. diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare/Pointers0.thy --- a/src/HOL/Hoare/Pointers0.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare/Pointers0.thy Fri Sep 20 19:51:08 2024 +0200 @@ -24,9 +24,9 @@ syntax "_fassign" :: "'a::ref => id => 'v => 's com" - ("(2_^._ :=/ _)" [70,1000,65] 61) + (\(2_^._ :=/ _)\ [70,1000,65] 61) "_faccess" :: "'a::ref => ('a::ref \ 'v) => 'v" - ("_^._" [65,1000] 65) + (\_^._\ [65,1000] 65) translations "p^.f := e" => "f := CONST fun_upd f p e" "p^.f" => "f p" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare/SchorrWaite.thy --- a/src/HOL/Hoare/SchorrWaite.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare/SchorrWaite.thy Fri Sep 20 19:51:08 2024 +0200 @@ -85,7 +85,7 @@ definition \ \Restriction of a relation\ - restr ::"('a \ 'a) set \ ('a \ bool) \ ('a \ 'a) set" ("(_/ | _)" [50, 51] 50) + restr ::"('a \ 'a) set \ ('a \ bool) \ ('a \ 'a) set" (\(_/ | _)\ [50, 51] 50) where "restr r m = {(x,y). (x,y) \ r \ \ m x}" text \Rewrite rules for the restriction of a relation\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare/Separation.thy --- a/src/HOL/Hoare/Separation.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare/Separation.thy Fri Sep 20 19:51:08 2024 +0200 @@ -19,7 +19,7 @@ text\The semantic definition of a few connectives:\ -definition ortho :: "heap \ heap \ bool" (infix "\" 55) +definition ortho :: "heap \ heap \ bool" (infix \\\ 55) where "h1 \ h2 \ dom h1 \ dom h2 = {}" definition is_empty :: "heap \ bool" @@ -51,10 +51,10 @@ bound Hs - otherwise they may bind the implicit H.\ syntax - "_emp" :: "bool" ("emp") - "_singl" :: "nat \ nat \ bool" ("[_ \ _]") - "_star" :: "bool \ bool \ bool" (infixl "**" 60) - "_wand" :: "bool \ bool \ bool" (infixl "-*" 60) + "_emp" :: "bool" (\emp\) + "_singl" :: "nat \ nat \ bool" (\[_ \ _]\) + "_star" :: "bool \ bool \ bool" (infixl \**\ 60) + "_wand" :: "bool \ bool \ bool" (infixl \-*\ 60) (* FIXME does not handle "_idtdummy" *) ML \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/OG_Hoare.thy --- a/src/HOL/Hoare_Parallel/OG_Hoare.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy Fri Sep 20 19:51:08 2024 +0200 @@ -34,8 +34,8 @@ interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) " inductive - oghoare :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\- _//_//_)" [90,55,90] 50) - and ann_hoare :: "'a ann_com \ 'a assn \ bool" ("(2\ _// _)" [60,90] 45) + oghoare :: "'a assn \ 'a com \ 'a assn \ bool" (\(3\- _//_//_)\ [90,55,90] 50) + and ann_hoare :: "'a ann_com \ 'a assn \ bool" (\(2\ _// _)\ [60,90] 45) where AnnBasic: "r \ {s. f s \ q} \ \ (AnnBasic r f) q" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/OG_Syntax.thy --- a/src/HOL/Hoare_Parallel/OG_Syntax.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/OG_Syntax.thy Fri Sep 20 19:51:08 2024 +0200 @@ -5,19 +5,19 @@ text\Syntax for commands and for assertions and boolean expressions in commands \com\ and annotated commands \ann_com\.\ -abbreviation Skip :: "'a com" ("SKIP" 63) +abbreviation Skip :: "'a com" (\SKIP\ 63) where "SKIP \ Basic id" -abbreviation AnnSkip :: "'a assn \ 'a ann_com" ("_//SKIP" [90] 63) +abbreviation AnnSkip :: "'a assn \ 'a ann_com" (\_//SKIP\ [90] 63) where "r SKIP \ AnnBasic r id" notation - Seq ("_,,/ _" [55, 56] 55) and - AnnSeq ("_;;/ _" [60,61] 60) + Seq (\_,,/ _\ [55, 56] 55) and + AnnSeq (\_;;/ _\ [60,61] 60) syntax - "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) - "_AnnAssign" :: "'a assn \ idt \ 'b \ 'a com" ("(_ \_ :=/ _)" [90,70,65] 61) + "_Assign" :: "idt \ 'b \ 'a com" (\(\_ :=/ _)\ [70, 65] 61) + "_AnnAssign" :: "'a assn \ idt \ 'b \ 'a com" (\(_ \_ :=/ _)\ [90,70,65] 61) translations "\x := a" \ "CONST Basic \\(_update_name x (\_. a))\" @@ -25,23 +25,23 @@ syntax "_AnnCond1" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com \ 'a ann_com" - ("_ //IF _ /THEN _ /ELSE _ /FI" [90,0,0,0] 61) + (\_ //IF _ /THEN _ /ELSE _ /FI\ [90,0,0,0] 61) "_AnnCond2" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com" - ("_ //IF _ /THEN _ /FI" [90,0,0] 61) + (\_ //IF _ /THEN _ /FI\ [90,0,0] 61) "_AnnWhile" :: "'a assn \ 'a bexp \ 'a assn \ 'a ann_com \ 'a ann_com" - ("_ //WHILE _ /INV _ //DO _//OD" [90,0,0,0] 61) + (\_ //WHILE _ /INV _ //DO _//OD\ [90,0,0,0] 61) "_AnnAwait" :: "'a assn \ 'a bexp \ 'a com \ 'a ann_com" - ("_ //AWAIT _ /THEN /_ /END" [90,0,0] 61) - "_AnnAtom" :: "'a assn \ 'a com \ 'a ann_com" ("_//\_\" [90,0] 61) - "_AnnWait" :: "'a assn \ 'a bexp \ 'a ann_com" ("_//WAIT _ END" [90,0] 61) + (\_ //AWAIT _ /THEN /_ /END\ [90,0,0] 61) + "_AnnAtom" :: "'a assn \ 'a com \ 'a ann_com" (\_//\_\\ [90,0] 61) + "_AnnWait" :: "'a assn \ 'a bexp \ 'a ann_com" (\_//WAIT _ END\ [90,0] 61) "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" - ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61) - "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("IF _ THEN _ FI" [0,0] 56) + (\(0IF _/ THEN _/ ELSE _/ FI)\ [0, 0, 0] 61) + "_Cond2" :: "'a bexp \ 'a com \ 'a com" (\IF _ THEN _ FI\ [0,0] 56) "_While_inv" :: "'a bexp \ 'a assn \ 'a com \ 'a com" - ("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61) + (\(0WHILE _/ INV _ //DO _ /OD)\ [0, 0, 0] 61) "_While" :: "'a bexp \ 'a com \ 'a com" - ("(0WHILE _ //DO _ /OD)" [0, 0] 61) + (\(0WHILE _ //DO _ /OD)\ [0, 0] 61) translations "IF b THEN c1 ELSE c2 FI" \ "CONST Cond \b\ c1 c2" @@ -59,12 +59,12 @@ nonterminal prgs syntax - "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" [57] 56) - "_prg" :: "['a, 'a] \ prgs" ("_//_" [60, 90] 57) - "_prgs" :: "['a, 'a, prgs] \ prgs" ("_//_//\//_" [60,90,57] 57) + "_PAR" :: "prgs \ 'a" (\COBEGIN//_//COEND\ [57] 56) + "_prg" :: "['a, 'a] \ prgs" (\_//_\ [60, 90] 57) + "_prgs" :: "['a, 'a, prgs] \ prgs" (\_//_//\//_\ [60,90,57] 57) "_prg_scheme" :: "['a, 'a, 'a, 'a, 'a] \ prgs" - ("SCHEME [_ \ _ < _] _// _" [0,0,0,60, 90] 57) + (\SCHEME [_ \ _ < _] _// _\ [0,0,0,60, 90] 57) translations "_prg c q" \ "[(CONST Some c, q)]" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/OG_Tactics.thy --- a/src/HOL/Hoare_Parallel/OG_Tactics.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy Fri Sep 20 19:51:08 2024 +0200 @@ -205,7 +205,7 @@ \ interfree (map (\k. (c k, Q k)) [a.. bool " ("[\] _" [0] 45) where +definition map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \ bool " (\[\] _\ [0] 45) where "[\] Ts == (\ic q. Ts!i=(Some c, q) \ \ c q)" lemma MapAnnEmpty: "[\] []" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/OG_Tran.thy --- a/src/HOL/Hoare_Parallel/OG_Tran.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/OG_Tran.thy Fri Sep 20 19:51:08 2024 +0200 @@ -20,11 +20,11 @@ ann_transition :: "(('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a)) set" and transition :: "(('a com \ 'a) \ ('a com \ 'a)) set" and ann_transition' :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" - ("_ -1\ _"[81,81] 100) + (\_ -1\ _\[81,81] 100) and transition' :: "('a com \ 'a) \ ('a com \ 'a) \ bool" - ("_ -P1\ _"[81,81] 100) + (\_ -P1\ _\[81,81] 100) and transitions :: "('a com \ 'a) \ ('a com \ 'a) \ bool" - ("_ -P*\ _"[81,81] 100) + (\_ -P*\ _\[81,81] 100) where "con_0 -1\ con_1 \ (con_0, con_1) \ ann_transition" | "con_0 -P1\ con_1 \ (con_0, con_1) \ transition" @@ -70,17 +70,17 @@ abbreviation ann_transition_n :: "('a ann_com_op \ 'a) \ nat \ ('a ann_com_op \ 'a) - \ bool" ("_ -_\ _"[81,81] 100) where + \ bool" (\_ -_\ _\[81,81] 100) where "con_0 -n\ con_1 \ (con_0, con_1) \ ann_transition ^^ n" abbreviation ann_transitions :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" - ("_ -*\ _"[81,81] 100) where + (\_ -*\ _\[81,81] 100) where "con_0 -*\ con_1 \ (con_0, con_1) \ ann_transition\<^sup>*" abbreviation transition_n :: "('a com \ 'a) \ nat \ ('a com \ 'a) \ bool" - ("_ -P_\ _"[81,81,81] 100) where + (\_ -P_\ _\[81,81,81] 100) where "con_0 -Pn\ con_1 \ (con_0, con_1) \ transition ^^ n" subsection \Definition of Semantics\ @@ -97,7 +97,7 @@ definition SEM :: "'a com \ 'a set \ 'a set" where "SEM c S \ \(sem c ` S)" -abbreviation Omega :: "'a com" ("\" 63) +abbreviation Omega :: "'a com" (\\\ 63) where "\ \ While UNIV UNIV (Basic id)" primrec fwhile :: "'a bexp \ 'a com \ nat \ 'a com" where @@ -294,10 +294,10 @@ section \Validity of Correctness Formulas\ -definition com_validity :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\= _// _//_)" [90,55,90] 50) where +definition com_validity :: "'a assn \ 'a com \ 'a assn \ bool" (\(3\= _// _//_)\ [90,55,90] 50) where "\= p c q \ SEM c p \ q" -definition ann_com_validity :: "'a ann_com \ 'a assn \ bool" ("\ _ _" [60,90] 45) where +definition ann_com_validity :: "'a ann_com \ 'a assn \ bool" (\\ _ _\ [60,90] 45) where "\ c q \ ann_SEM c (pre c) \ q" end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/Quote_Antiquote.thy --- a/src/HOL/Hoare_Parallel/Quote_Antiquote.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/Quote_Antiquote.thy Fri Sep 20 19:51:08 2024 +0200 @@ -3,9 +3,9 @@ theory Quote_Antiquote imports Main begin syntax - "_quote" :: "'b \ ('a \ 'b)" ("(\_\)" [0] 1000) - "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) - "_Assert" :: "'a \ 'a set" ("(\_\)" [0] 1000) + "_quote" :: "'b \ ('a \ 'b)" (\(\_\)\ [0] 1000) + "_antiquote" :: "('a \ 'b) \ 'b" (\\_\ [1000] 1000) + "_Assert" :: "'a \ 'a set" (\(\_\)\ [0] 1000) translations "\b\" \ "CONST Collect \b\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/RG_Hoare.thy --- a/src/HOL/Hoare_Parallel/RG_Hoare.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ inductive rghoare :: "['a com, 'a set, ('a \ 'a) set, ('a \ 'a) set, 'a set] \ bool" - ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) + (\\ _ sat [_, _, _, _]\ [60,0,0,0,0] 45) where Basic: "\ pre \ {s. f s \ post}; {(s,t). s \ pre \ (t=f s \ t=s)} \ guar; stable pre rely; stable post rely \ @@ -59,7 +59,7 @@ inductive par_rghoare :: "('a rgformula) list \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" - ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) + (\\ _ SAT [_, _, _, _]\ [60,0,0,0,0] 45) where Parallel: "\ \i (\j\{j. j j\i}. Guar(xs!j)) \ Rely(xs!i); diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/RG_Syntax.thy --- a/src/HOL/Hoare_Parallel/RG_Syntax.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/RG_Syntax.thy Fri Sep 20 19:51:08 2024 +0200 @@ -4,19 +4,19 @@ imports RG_Hoare Quote_Antiquote begin -abbreviation Skip :: "'a com" ("SKIP") +abbreviation Skip :: "'a com" (\SKIP\) where "SKIP \ Basic id" -notation Seq ("(_;;/ _)" [60,61] 60) +notation Seq (\(_;;/ _)\ [60,61] 60) syntax - "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) - "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61) - "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("(0IF _ THEN _ FI)" [0,0] 56) - "_While" :: "'a bexp \ 'a com \ 'a com" ("(0WHILE _ /DO _ /OD)" [0, 0] 61) - "_Await" :: "'a bexp \ 'a com \ 'a com" ("(0AWAIT _ /THEN /_ /END)" [0,0] 61) - "_Atom" :: "'a com \ 'a com" ("(\_\)" 61) - "_Wait" :: "'a bexp \ 'a com" ("(0WAIT _ END)" 61) + "_Assign" :: "idt \ 'b \ 'a com" (\(\_ :=/ _)\ [70, 65] 61) + "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" (\(0IF _/ THEN _/ ELSE _/FI)\ [0, 0, 0] 61) + "_Cond2" :: "'a bexp \ 'a com \ 'a com" (\(0IF _ THEN _ FI)\ [0,0] 56) + "_While" :: "'a bexp \ 'a com \ 'a com" (\(0WHILE _ /DO _ /OD)\ [0, 0] 61) + "_Await" :: "'a bexp \ 'a com \ 'a com" (\(0AWAIT _ /THEN /_ /END)\ [0,0] 61) + "_Atom" :: "'a com \ 'a com" (\(\_\)\ 61) + "_Wait" :: "'a bexp \ 'a com" (\(0WAIT _ END)\ 61) translations "\x := a" \ "CONST Basic \\(_update_name x (\_. a))\" @@ -30,9 +30,9 @@ nonterminal prgs syntax - "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" 60) - "_prg" :: "'a \ prgs" ("_" 57) - "_prgs" :: "['a, prgs] \ prgs" ("_//\//_" [60,57] 57) + "_PAR" :: "prgs \ 'a" (\COBEGIN//_//COEND\ 60) + "_prg" :: "'a \ prgs" (\_\ 57) + "_prgs" :: "['a, prgs] \ prgs" (\_//\//_\ [60,57] 57) translations "_prg a" \ "[a]" @@ -40,7 +40,7 @@ "_PAR ps" \ "ps" syntax - "_prg_scheme" :: "['a, 'a, 'a, 'a] \ prgs" ("SCHEME [_ \ _ < _] _" [0,0,0,60] 57) + "_prg_scheme" :: "['a, 'a, 'a, 'a] \ prgs" (\SCHEME [_ \ _ < _] _\ [0,0,0,60] 57) translations "_prg_scheme j i k c" \ "(CONST map (\i. c) [j..Translations for variables before and after a transition:\ syntax - "_before" :: "id \ 'a" ("\_") - "_after" :: "id \ 'a" ("\_") + "_before" :: "id \ 'a" (\\_\) + "_after" :: "id \ 'a" (\\_\) translations "\x" \ "x \CONST fst" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Hoare_Parallel/RG_Tran.thy --- a/src/HOL/Hoare_Parallel/RG_Tran.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ inductive_set etran :: "('a conf \ 'a conf) set" - and etran' :: "'a conf \ 'a conf \ bool" ("_ -e\ _" [81,81] 80) + and etran' :: "'a conf \ 'a conf \ bool" (\_ -e\ _\ [81,81] 80) where "P -e\ Q \ (P,Q) \ etran" | Env: "(P, s) -e\ (P, t)" @@ -24,8 +24,8 @@ inductive_set ctran :: "('a conf \ 'a conf) set" - and ctran' :: "'a conf \ 'a conf \ bool" ("_ -c\ _" [81,81] 80) - and ctrans :: "'a conf \ 'a conf \ bool" ("_ -c*\ _" [81,81] 80) + and ctran' :: "'a conf \ 'a conf \ bool" (\_ -c\ _\ [81,81] 80) + and ctrans :: "'a conf \ 'a conf \ bool" (\_ -c*\ _\ [81,81] 80) where "P -c\ Q \ (P,Q) \ ctran" | "P -c*\ Q \ (P,Q) \ ctran\<^sup>*" @@ -52,14 +52,14 @@ inductive_set par_etran :: "('a par_conf \ 'a par_conf) set" - and par_etran' :: "['a par_conf,'a par_conf] \ bool" ("_ -pe\ _" [81,81] 80) + and par_etran' :: "['a par_conf,'a par_conf] \ bool" (\_ -pe\ _\ [81,81] 80) where "P -pe\ Q \ (P,Q) \ par_etran" | ParEnv: "(Ps, s) -pe\ (Ps, t)" inductive_set par_ctran :: "('a par_conf \ 'a par_conf) set" - and par_ctran' :: "['a par_conf,'a par_conf] \ bool" ("_ -pc\ _" [81,81] 80) + and par_ctran' :: "['a par_conf,'a par_conf] \ bool" (\_ -pc\ _\ [81,81] 80) where "P -pc\ Q \ (P,Q) \ par_ctran" | ParComp: "\i (r, t)\ \ (Ps, s) -pc\ (Ps[i:=r], t)" @@ -367,7 +367,7 @@ (fst (last c) = None \ snd (last c) \ post)}" definition com_validity :: "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" - ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) where + (\\ _ sat [_, _, _, _]\ [60,0,0,0,0] 45) where "\ P sat [pre, rely, guar, post] \ \s. cp (Some P) s \ assum(pre, rely) \ comm(guar, post)" @@ -386,7 +386,7 @@ (All_None (fst (last c)) \ snd( last c) \ post)}" definition par_com_validity :: "'a par_com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set -\ 'a set \ bool" ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) where +\ 'a set \ bool" (\\ _ SAT [_, _, _, _]\ [60,0,0,0,0] 45) where "\ Ps SAT [pre, rely, guar, post] \ \s. par_cp Ps s \ par_assum(pre, rely) \ par_comm(guar, post)" @@ -409,7 +409,7 @@ (\li \ (clist!l)!j -e\ (clist!l)! Suc j))) \ (c!j -pe\ c!Suc j \ (\i (clist!i)! Suc j)))" -definition conjoin :: "'a par_confs \ ('a confs) list \ bool" ("_ \ _" [65,65] 64) where +definition conjoin :: "'a par_confs \ ('a confs) list \ bool" (\_ \ _\ [65,65] 64) where "c \ clist \ (same_length c clist) \ (same_state c clist) \ (same_program c clist) \ (compat_label c clist)" subsubsection \Some previous lemmas\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/ACom.thy --- a/src/HOL/IMP/ACom.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/ACom.thy Fri Sep 20 19:51:08 2024 +0200 @@ -7,15 +7,15 @@ begin datatype 'a acom = - SKIP 'a ("SKIP {_}" 61) | - Assign vname aexp 'a ("(_ ::= _/ {_})" [1000, 61, 0] 61) | - Seq "('a acom)" "('a acom)" ("_;;//_" [60, 61] 60) | + SKIP 'a (\SKIP {_}\ 61) | + Assign vname aexp 'a (\(_ ::= _/ {_})\ [1000, 61, 0] 61) | + Seq "('a acom)" "('a acom)" (\_;;//_\ [60, 61] 60) | If bexp 'a "('a acom)" 'a "('a acom)" 'a - ("(IF _/ THEN ({_}/ _)/ ELSE ({_}/ _)//{_})" [0, 0, 0, 61, 0, 0] 61) | + (\(IF _/ THEN ({_}/ _)/ ELSE ({_}/ _)//{_})\ [0, 0, 0, 61, 0, 0] 61) | While 'a bexp 'a "('a acom)" 'a - ("({_}//WHILE _//DO ({_}//_)//{_})" [0, 0, 0, 61, 0] 61) + (\({_}//WHILE _//DO ({_}//_)//{_})\ [0, 0, 0, 61, 0] 61) -notation com.SKIP ("SKIP") +notation com.SKIP (\SKIP\) text_raw\\snip{stripdef}{1}{1}{%\ fun strip :: "'a acom \ com" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/AExp.thy --- a/src/HOL/IMP/AExp.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/AExp.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,10 +27,10 @@ text \A little syntax magic to write larger states compactly:\ -definition null_state ("<>") where +definition null_state (\<>\) where "null_state \ \x. 0" syntax - "_State" :: "updbinds => 'a" ("<_>") + "_State" :: "updbinds => 'a" (\<_>\) translations "_State ms" == "_Update <> ms" "_State (_updbinds b bs)" <= "_Update (_State b) bs" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Abs_Int0.thy --- a/src/HOL/IMP/Abs_Int0.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Abs_Int0.thy Fri Sep 20 19:51:08 2024 +0200 @@ -317,20 +317,20 @@ assumes h: "m x \ h" begin -definition m_s :: "'av st \ vname set \ nat" ("m\<^sub>s") where +definition m_s :: "'av st \ vname set \ nat" (\m\<^sub>s\) where "m_s S X = (\ x \ X. m(S x))" lemma m_s_h: "finite X \ m_s S X \ h * card X" by(simp add: m_s_def) (metis mult.commute of_nat_id sum_bounded_above[OF h]) -fun m_o :: "'av st option \ vname set \ nat" ("m\<^sub>o") where +fun m_o :: "'av st option \ vname set \ nat" (\m\<^sub>o\) where "m_o (Some S) X = m_s S X" | "m_o None X = h * card X + 1" lemma m_o_h: "finite X \ m_o opt X \ (h*card X + 1)" by(cases opt)(auto simp add: m_s_h le_SucI dest: m_s_h) -definition m_c :: "'av st option acom \ nat" ("m\<^sub>c") where +definition m_c :: "'av st option acom \ nat" (\m\<^sub>c\) where "m_c C = sum_list (map (\a. m_o a (vars C)) (annos C))" text\Upper complexity bound:\ @@ -359,14 +359,14 @@ the finitely many variables in the program change. That the others do not change follows because they remain \<^term>\\\.\ -fun top_on_st :: "'av st \ vname set \ bool" ("top'_on\<^sub>s") where +fun top_on_st :: "'av st \ vname set \ bool" (\top'_on\<^sub>s\) where "top_on_st S X = (\x\X. S x = \)" -fun top_on_opt :: "'av st option \ vname set \ bool" ("top'_on\<^sub>o") where +fun top_on_opt :: "'av st option \ vname set \ bool" (\top'_on\<^sub>o\) where "top_on_opt (Some S) X = top_on_st S X" | "top_on_opt None X = True" -definition top_on_acom :: "'av st option acom \ vname set \ bool" ("top'_on\<^sub>c") where +definition top_on_acom :: "'av st option acom \ vname set \ bool" (\top'_on\<^sub>c\) where "top_on_acom C X = (\a \ set(annos C). top_on_opt a X)" lemma top_on_top: "top_on_opt \ X" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Abs_Int1.thy --- a/src/HOL/IMP/Abs_Int1.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Abs_Int1.thy Fri Sep 20 19:51:08 2024 +0200 @@ -105,19 +105,19 @@ assumes h: "m x \ h" begin -definition m_s :: "'av st \ vname set \ nat" ("m\<^sub>s") where +definition m_s :: "'av st \ vname set \ nat" (\m\<^sub>s\) where "m_s S X = (\ x \ X. m(fun S x))" lemma m_s_h: "finite X \ m_s S X \ h * card X" by(simp add: m_s_def) (metis mult.commute of_nat_id sum_bounded_above[OF h]) -definition m_o :: "'av st option \ vname set \ nat" ("m\<^sub>o") where +definition m_o :: "'av st option \ vname set \ nat" (\m\<^sub>o\) where "m_o opt X = (case opt of None \ h * card X + 1 | Some S \ m_s S X)" lemma m_o_h: "finite X \ m_o opt X \ (h*card X + 1)" by(auto simp add: m_o_def m_s_h le_SucI split: option.split dest:m_s_h) -definition m_c :: "'av st option acom \ nat" ("m\<^sub>c") where +definition m_c :: "'av st option acom \ nat" (\m\<^sub>c\) where "m_c C = sum_list (map (\a. m_o a (vars C)) (annos C))" text\Upper complexity bound:\ @@ -134,14 +134,14 @@ end -fun top_on_st :: "'a::order_top st \ vname set \ bool" ("top'_on\<^sub>s") where +fun top_on_st :: "'a::order_top st \ vname set \ bool" (\top'_on\<^sub>s\) where "top_on_st S X = (\x\X. fun S x = \)" -fun top_on_opt :: "'a::order_top st option \ vname set \ bool" ("top'_on\<^sub>o") where +fun top_on_opt :: "'a::order_top st option \ vname set \ bool" (\top'_on\<^sub>o\) where "top_on_opt (Some S) X = top_on_st S X" | "top_on_opt None X = True" -definition top_on_acom :: "'a::order_top st option acom \ vname set \ bool" ("top'_on\<^sub>c") where +definition top_on_acom :: "'a::order_top st option acom \ vname set \ bool" (\top'_on\<^sub>c\) where "top_on_acom C X = (\a \ set(annos C). top_on_opt a X)" lemma top_on_top: "top_on_opt (\::_ st option) X" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Abs_Int2_ivl.thy --- a/src/HOL/IMP/Abs_Int2_ivl.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Abs_Int2_ivl.thy Fri Sep 20 19:51:08 2024 +0200 @@ -21,7 +21,7 @@ quotient_type ivl = eint2 / eq_ivl by(rule equivpI)(auto simp: reflp_def symp_def transp_def eq_ivl_def) -abbreviation ivl_abbr :: "eint \ eint \ ivl" ("[_, _]") where +abbreviation ivl_abbr :: "eint \ eint \ ivl" (\[_, _]\) where "[l,h] == abs_ivl(l,h)" lift_definition \_ivl :: "ivl \ int set" is \_rep diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Abs_Int3.thy --- a/src/HOL/IMP/Abs_Int3.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Abs_Int3.thy Fri Sep 20 19:51:08 2024 +0200 @@ -7,10 +7,10 @@ begin class widen = -fixes widen :: "'a \ 'a \ 'a" (infix "\" 65) +fixes widen :: "'a \ 'a \ 'a" (infix \\\ 65) class narrow = -fixes narrow :: "'a \ 'a \ 'a" (infix "\" 65) +fixes narrow :: "'a \ 'a \ 'a" (infix \\\ 65) class wn = widen + narrow + order + assumes widen1: "x \ x \ y" @@ -398,7 +398,7 @@ done -definition n_s :: "'av st \ vname set \ nat" ("n\<^sub>s") where +definition n_s :: "'av st \ vname set \ nat" (\n\<^sub>s\) where "n\<^sub>s S X = (\x\X. n(fun S x))" lemma n_s_narrow_rep: @@ -422,7 +422,7 @@ apply(auto simp add: less_eq_st_rep_iff eq_st_def fun_eq_iff n_s_narrow_rep) done -definition n_o :: "'av st option \ vname set \ nat" ("n\<^sub>o") where +definition n_o :: "'av st option \ vname set \ nat" (\n\<^sub>o\) where "n\<^sub>o opt X = (case opt of None \ 0 | Some S \ n\<^sub>s S X + 1)" lemma n_o_narrow: @@ -433,7 +433,7 @@ done -definition n_c :: "'av st option acom \ nat" ("n\<^sub>c") where +definition n_c :: "'av st option acom \ nat" (\n\<^sub>c\) where "n\<^sub>c C = sum_list (map (\a. n\<^sub>o a (vars C)) (annos C))" lemma less_annos_iff: "(C1 < C2) = (C1 \ C2 \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Big_Step.thy --- a/src/HOL/IMP/Big_Step.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Big_Step.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ text_raw\\snip{BigStepdef}{0}{1}{%\ inductive - big_step :: "com \ state \ state \ bool" (infix "\" 55) + big_step :: "com \ state \ state \ bool" (infix \\\ 55) where Skip: "(SKIP,s) \ s" | Assign: "(x ::= a,s) \ s(x := aval a s)" | @@ -163,7 +163,7 @@ \ text_raw\\snip{BigStepEquiv}{0}{1}{%\ abbreviation - equiv_c :: "com \ com \ bool" (infix "\" 50) where + equiv_c :: "com \ com \ bool" (infix \\\ 50) where "c \ c' \ (\s t. (c,s) \ t = (c',s) \ t)" text_raw\}%endsnip\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/C_like.thy --- a/src/HOL/IMP/C_like.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/C_like.thy Fri Sep 20 19:51:08 2024 +0200 @@ -4,7 +4,7 @@ type_synonym state = "nat \ nat" -datatype aexp = N nat | Deref aexp ("!") | Plus aexp aexp +datatype aexp = N nat | Deref aexp (\!\) | Plus aexp aexp fun aval :: "aexp \ state \ nat" where "aval (N n) s = n" | @@ -22,14 +22,14 @@ datatype com = SKIP - | Assign aexp aexp ("_ ::= _" [61, 61] 61) + | Assign aexp aexp (\_ ::= _\ [61, 61] 61) | New aexp aexp - | Seq com com ("_;/ _" [60, 61] 60) - | If bexp com com ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) - | While bexp com ("(WHILE _/ DO _)" [0, 61] 61) + | Seq com com (\_;/ _\ [60, 61] 60) + | If bexp com com (\(IF _/ THEN _/ ELSE _)\ [0, 0, 61] 61) + | While bexp com (\(WHILE _/ DO _)\ [0, 61] 61) inductive - big_step :: "com \ state \ nat \ state \ nat \ bool" (infix "\" 55) + big_step :: "com \ state \ nat \ state \ nat \ bool" (infix \\\ 55) where Skip: "(SKIP,sn) \ sn" | Assign: "(lhs ::= a,s,n) \ (s(aval lhs s := aval a s),n)" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Collecting.thy --- a/src/HOL/IMP/Collecting.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Collecting.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,10 +9,10 @@ subsubsection "The generic Step function" notation - sup (infixl "\" 65) and - inf (infixl "\" 70) and - bot ("\") and - top ("\") + sup (infixl \\\ 65) and + inf (infixl \\\ 70) and + bot (\\\) and + top (\\\) context fixes f :: "vname \ aexp \ 'a \ 'a::sup" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Com.thy --- a/src/HOL/IMP/Com.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Com.thy Fri Sep 20 19:51:08 2024 +0200 @@ -4,9 +4,9 @@ datatype com = SKIP - | Assign vname aexp ("_ ::= _" [1000, 61] 61) - | Seq com com ("_;;/ _" [60, 61] 60) - | If bexp com com ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) - | While bexp com ("(WHILE _/ DO _)" [0, 61] 61) + | Assign vname aexp (\_ ::= _\ [1000, 61] 61) + | Seq com com (\_;;/ _\ [60, 61] 60) + | If bexp com com (\(IF _/ THEN _/ ELSE _)\ [0, 0, 61] 61) + | While bexp com (\(WHILE _/ DO _)\ [0, 61] 61) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Compiler.thy --- a/src/HOL/IMP/Compiler.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Compiler.thy Fri Sep 20 19:51:08 2024 +0200 @@ -20,7 +20,7 @@ Similarly, we will want to access the ith element of a list, where \<^term>\i\ is an \<^typ>\int\. \ -fun inth :: "'a list \ int \ 'a" (infixl "!!" 100) where +fun inth :: "'a list \ int \ 'a" (infixl \!!\ 100) where "(x # xs) !! i = (if i = 0 then x else xs !! (i - 1))" text \ @@ -37,7 +37,7 @@ abbreviation (output) "isize xs == int (length xs)" -notation isize ("size") +notation isize (\size\) subsection "Instructions and Stack Machine" @@ -66,7 +66,7 @@ definition exec1 :: "instr list \ config \ config \ bool" - ("(_/ \ (_ \/ _))" [59,0,59] 60) + (\(_/ \ (_ \/ _))\ [59,0,59] 60) where "P \ c \ c' = (\i s stk. c = (i,s,stk) \ c' = iexec(P!!i) (i,s,stk) \ 0 \ i \ i < size P)" @@ -77,7 +77,7 @@ by (simp add: exec1_def) abbreviation - exec :: "instr list \ config \ config \ bool" ("(_/ \ (_ \*/ _))" 50) + exec :: "instr list \ config \ config \ bool" (\(_/ \ (_ \*/ _))\ 50) where "exec P \ star (exec1 P)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Compiler2.thy --- a/src/HOL/IMP/Compiler2.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Compiler2.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ text \Execution in \<^term>\n\ steps for simpler induction\ primrec exec_n :: "instr list \ config \ nat \ config \ bool" - ("_/ \ (_ \^_/ _)" [65,0,1000,55] 55) + (\_/ \ (_ \^_/ _)\ [65,0,1000,55] 55) where "P \ c \^0 c' = (c'=c)" | "P \ c \^(Suc n) c'' = (\c'. (P \ c \ c') \ P \ c' \^n c'')" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Def_Init_Big.thy --- a/src/HOL/IMP/Def_Init_Big.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Def_Init_Big.thy Fri Sep 20 19:51:08 2024 +0200 @@ -7,7 +7,7 @@ subsection "Initialization-Sensitive Big Step Semantics" inductive - big_step :: "(com \ state option) \ state option \ bool" (infix "\" 55) + big_step :: "(com \ state option) \ state option \ bool" (infix \\\ 55) where None: "(c,None) \ None" | Skip: "(SKIP,s) \ s" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Def_Init_Small.thy --- a/src/HOL/IMP/Def_Init_Small.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Def_Init_Small.thy Fri Sep 20 19:51:08 2024 +0200 @@ -7,7 +7,7 @@ subsection "Initialization-Sensitive Small Step Semantics" inductive - small_step :: "(com \ state) \ (com \ state) \ bool" (infix "\" 55) + small_step :: "(com \ state) \ (com \ state) \ bool" (infix \\\ 55) where Assign: "aval a s = Some i \ (x ::= a, s) \ (SKIP, s(x := Some i))" | @@ -21,7 +21,7 @@ lemmas small_step_induct = small_step.induct[split_format(complete)] -abbreviation small_steps :: "com * state \ com * state \ bool" (infix "\*" 55) +abbreviation small_steps :: "com * state \ com * state \ bool" (infix \\*\ 55) where "x \* y == star small_step x y" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Finite_Reachable.thy --- a/src/HOL/IMP/Finite_Reachable.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Finite_Reachable.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,7 +15,7 @@ text\Proofs need induction on the length of a small-step reduction sequence.\ fun small_stepsn :: "com * state \ nat \ com * state \ bool" - ("_ \'(_') _" [55,0,55] 55) where + (\_ \'(_') _\ [55,0,55] 55) where "(cs \(0) cs') = (cs' = cs)" | "cs \(Suc n) cs'' = (\cs'. cs \ cs' \ cs' \(n) cs'')" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Hoare.thy --- a/src/HOL/IMP/Hoare.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Hoare.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,15 +9,15 @@ type_synonym assn = "state \ bool" definition -hoare_valid :: "assn \ com \ assn \ bool" ("\ {(1_)}/ (_)/ {(1_)}" 50) where +hoare_valid :: "assn \ com \ assn \ bool" (\\ {(1_)}/ (_)/ {(1_)}\ 50) where "\ {P}c{Q} = (\s t. P s \ (c,s) \ t \ Q t)" abbreviation state_subst :: "state \ aexp \ vname \ state" - ("_[_'/_]" [1000,0,0] 999) + (\_[_'/_]\ [1000,0,0] 999) where "s[a/x] == s(x := aval a s)" inductive - hoare :: "assn \ com \ assn \ bool" ("\ ({(1_)}/ (_)/ {(1_)})" 50) + hoare :: "assn \ com \ assn \ bool" (\\ ({(1_)}/ (_)/ {(1_)})\ 50) where Skip: "\ {P} SKIP {P}" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Hoare_Total.thy --- a/src/HOL/IMP/Hoare_Total.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Hoare_Total.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ works if execution is deterministic (which it is in our case).\ definition hoare_tvalid :: "assn \ com \ assn \ bool" - ("\\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where + (\\\<^sub>t {(1_)}/ (_)/ {(1_)}\ 50) where "\\<^sub>t {P}c{Q} \ (\s. P s \ (\t. (c,s) \ t \ Q t))" text\Provability of Hoare triples in the proof system for total @@ -22,7 +22,7 @@ \<^term>\While\-rule.\ inductive - hoaret :: "assn \ com \ assn \ bool" ("\\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50) + hoaret :: "assn \ com \ assn \ bool" (\\\<^sub>t ({(1_)}/ (_)/ {(1_)})\ 50) where Skip: "\\<^sub>t {P} SKIP {P}" | @@ -136,7 +136,7 @@ correctness. First we have to strengthen our notion of weakest precondition to take termination into account:\ -definition wpt :: "com \ assn \ assn" ("wp\<^sub>t") where +definition wpt :: "com \ assn \ assn" (\wp\<^sub>t\) where "wp\<^sub>t c Q = (\s. \t. (c,s) \ t \ Q t)" lemma [simp]: "wp\<^sub>t SKIP Q = Q" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Hoare_Total_EX.thy --- a/src/HOL/IMP/Hoare_Total_EX.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Hoare_Total_EX.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,11 +13,11 @@ to apply in program proofs.\ definition hoare_tvalid :: "assn \ com \ assn \ bool" - ("\\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where + (\\\<^sub>t {(1_)}/ (_)/ {(1_)}\ 50) where "\\<^sub>t {P}c{Q} \ (\s. P s \ (\t. (c,s) \ t \ Q t))" inductive - hoaret :: "assn \ com \ assn \ bool" ("\\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50) + hoaret :: "assn \ com \ assn \ bool" (\\\<^sub>t ({(1_)}/ (_)/ {(1_)})\ 50) where Skip: "\\<^sub>t {P} SKIP {P}" | @@ -68,7 +68,7 @@ qed fastforce+ -definition wpt :: "com \ assn \ assn" ("wp\<^sub>t") where +definition wpt :: "com \ assn \ assn" (\wp\<^sub>t\) where "wp\<^sub>t c Q = (\s. \t. (c,s) \ t \ Q t)" lemma [simp]: "wp\<^sub>t SKIP Q = Q" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Hoare_Total_EX2.thy --- a/src/HOL/IMP/Hoare_Total_EX2.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Hoare_Total_EX2.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,11 +17,11 @@ type_synonym assn2 = "(lvname \ nat) \ state \ bool" definition hoare_tvalid :: "assn2 \ com \ assn2 \ bool" - ("\\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where + (\\\<^sub>t {(1_)}/ (_)/ {(1_)}\ 50) where "\\<^sub>t {P}c{Q} \ (\l s. P l s \ (\t. (c,s) \ t \ Q l t))" inductive - hoaret :: "assn2 \ com \ assn2 \ bool" ("\\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50) + hoaret :: "assn2 \ com \ assn2 \ bool" (\\\<^sub>t ({(1_)}/ (_)/ {(1_)})\ 50) where Skip: "\\<^sub>t {P} SKIP {P}" | @@ -75,7 +75,7 @@ qed fastforce+ -definition wpt :: "com \ assn2 \ assn2" ("wp\<^sub>t") where +definition wpt :: "com \ assn2 \ assn2" (\wp\<^sub>t\) where "wp\<^sub>t c Q = (\l s. \t. (c,s) \ t \ Q l t)" lemma [simp]: "wp\<^sub>t SKIP Q = Q" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/OO.thy --- a/src/HOL/IMP/OO.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/OO.thy Fri Sep 20 19:51:08 2024 +0200 @@ -4,7 +4,7 @@ (* FIXME: move to HOL/Fun *) abbreviation fun_upd2 :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c \ 'a \ 'b \ 'c" - ("_/'((2_,_ :=/ _)')" [1000,0,0,0] 900) + (\_/'((2_,_ :=/ _)')\ [1000,0,0,0] 900) where "f(x,y := z) == f(x := (f x)(y := z))" type_synonym addr = nat @@ -18,12 +18,12 @@ Null | New | V string | - Faccess exp string ("_\/_" [63,1000] 63) | - Vassign string exp ("(_ ::=/ _)" [1000,61] 62) | - Fassign exp string exp ("(_\_ ::=/ _)" [63,0,62] 62) | - Mcall exp string exp ("(_\/_<_>)" [63,0,0] 63) | - Seq exp exp ("_;/ _" [61,60] 60) | - If bexp exp exp ("IF _/ THEN (2_)/ ELSE (2_)" [0,0,61] 61) + Faccess exp string (\_\/_\ [63,1000] 63) | + Vassign string exp (\(_ ::=/ _)\ [1000,61] 62) | + Fassign exp string exp (\(_\_ ::=/ _)\ [63,0,62] 62) | + Mcall exp string exp (\(_\/_<_>)\ [63,0,0] 63) | + Seq exp exp (\_;/ _\ [61,60] 60) | + If bexp exp exp (\IF _/ THEN (2_)/ ELSE (2_)\ [0,0,61] 61) and bexp = B bool | Not bexp | And bexp bexp | Eq exp exp type_synonym menv = "string \ exp" @@ -31,9 +31,9 @@ inductive big_step :: "menv \ exp \ config \ ref \ config \ bool" - ("(_ \/ (_/ \ _))" [60,0,60] 55) and + (\(_ \/ (_/ \ _))\ [60,0,60] 55) and bval :: "menv \ bexp \ config \ bool \ config \ bool" - ("_ \ _ \ _" [60,0,60] 55) + (\_ \ _ \ _\ [60,0,60] 55) where Null: "me \ (Null,c) \ (null,c)" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Poly_Types.thy --- a/src/HOL/IMP/Poly_Types.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Poly_Types.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,21 +9,21 @@ type_synonym tyenv = "vname \ ty" inductive atyping :: "tyenv \ aexp \ ty \ bool" - ("(1_/ \p/ (_ :/ _))" [50,0,50] 50) + (\(1_/ \p/ (_ :/ _))\ [50,0,50] 50) where "\ \p Ic i : Ity" | "\ \p Rc r : Rty" | "\ \p V x : \ x" | "\ \p a1 : \ \ \ \p a2 : \ \ \ \p Plus a1 a2 : \" -inductive btyping :: "tyenv \ bexp \ bool" (infix "\p" 50) +inductive btyping :: "tyenv \ bexp \ bool" (infix \\p\ 50) where "\ \p Bc v" | "\ \p b \ \ \p Not b" | "\ \p b1 \ \ \p b2 \ \ \p And b1 b2" | "\ \p a1 : \ \ \ \p a2 : \ \ \ \p Less a1 a2" -inductive ctyping :: "tyenv \ com \ bool" (infix "\p" 50) where +inductive ctyping :: "tyenv \ com \ bool" (infix \\p\ 50) where "\ \p SKIP" | "\ \p a : \(x) \ \ \p x ::= a" | "\ \p c1 \ \ \p c2 \ \ \p c1;;c2" | @@ -34,7 +34,7 @@ "type (Iv i) = Ity" | "type (Rv r) = Rty" -definition styping :: "tyenv \ state \ bool" (infix "\p" 50) +definition styping :: "tyenv \ state \ bool" (infix \\p\ 50) where "\ \p s \ (\x. type (s x) = \ x)" fun tsubst :: "(nat \ ty) \ ty \ ty" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Procs.thy --- a/src/HOL/IMP/Procs.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Procs.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,12 +8,12 @@ datatype com = SKIP - | Assign vname aexp ("_ ::= _" [1000, 61] 61) - | Seq com com ("_;;/ _" [60, 61] 60) - | If bexp com com ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) - | While bexp com ("(WHILE _/ DO _)" [0, 61] 61) - | Var vname com ("(1{VAR _;/ _})") - | Proc pname com com ("(1{PROC _ = _;/ _})") + | Assign vname aexp (\_ ::= _\ [1000, 61] 61) + | Seq com com (\_;;/ _\ [60, 61] 60) + | If bexp com com (\(IF _/ THEN _/ ELSE _)\ [0, 0, 61] 61) + | While bexp com (\(WHILE _/ DO _)\ [0, 61] 61) + | Var vname com (\(1{VAR _;/ _})\) + | Proc pname com com (\(1{PROC _ = _;/ _})\) | CALL pname definition "test_com = diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Procs_Dyn_Vars_Dyn.thy --- a/src/HOL/IMP/Procs_Dyn_Vars_Dyn.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Procs_Dyn_Vars_Dyn.thy Fri Sep 20 19:51:08 2024 +0200 @@ -6,7 +6,7 @@ type_synonym penv = "pname \ com" inductive - big_step :: "penv \ com \ state \ state \ bool" ("_ \ _ \ _" [60,0,60] 55) + big_step :: "penv \ com \ state \ state \ bool" (\_ \ _ \ _\ [60,0,60] 55) where Skip: "pe \ (SKIP,s) \ s" | Assign: "pe \ (x ::= a,s) \ s(x := aval a s)" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Procs_Stat_Vars_Dyn.thy --- a/src/HOL/IMP/Procs_Stat_Vars_Dyn.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Procs_Stat_Vars_Dyn.thy Fri Sep 20 19:51:08 2024 +0200 @@ -6,7 +6,7 @@ type_synonym penv = "(pname \ com) list" inductive - big_step :: "penv \ com \ state \ state \ bool" ("_ \ _ \ _" [60,0,60] 55) + big_step :: "penv \ com \ state \ state \ bool" (\_ \ _ \ _\ [60,0,60] 55) where Skip: "pe \ (SKIP,s) \ s" | Assign: "pe \ (x ::= a,s) \ s(x := aval a s)" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Procs_Stat_Vars_Stat.thy --- a/src/HOL/IMP/Procs_Stat_Vars_Stat.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Procs_Stat_Vars_Stat.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ inductive big_step :: "penv \ venv \ nat \ com \ store \ store \ bool" - ("_ \ _ \ _" [60,0,60] 55) + (\_ \ _ \ _\ [60,0,60] 55) where Skip: "e \ (SKIP,s) \ s" | Assign: "(pe,ve,f) \ (x ::= a,s) \ s(ve x := aval a (s o ve))" | diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Sec_Type_Expr.thy --- a/src/HOL/IMP/Sec_Type_Expr.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Sec_Type_Expr.thy Fri Sep 20 19:51:08 2024 +0200 @@ -50,11 +50,11 @@ abbreviation eq_le :: "state \ state \ level \ bool" - ("(_ = _ '(\ _'))" [51,51,0] 50) where + (\(_ = _ '(\ _'))\ [51,51,0] 50) where "s = s' (\ l) == (\ x. sec x \ l \ s x = s' x)" abbreviation eq_less :: "state \ state \ level \ bool" - ("(_ = _ '(< _'))" [51,51,0] 50) where + (\(_ = _ '(< _'))\ [51,51,0] 50) where "s = s' (< l) == (\ x. sec x < l \ s x = s' x)" lemma aval_eq_if_eq_le: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Sec_Typing.thy --- a/src/HOL/IMP/Sec_Typing.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Sec_Typing.thy Fri Sep 20 19:51:08 2024 +0200 @@ -7,7 +7,7 @@ subsubsection "Syntax Directed Typing" -inductive sec_type :: "nat \ com \ bool" ("(_/ \ _)" [0,0] 50) where +inductive sec_type :: "nat \ com \ bool" (\(_/ \ _)\ [0,0] 50) where Skip: "l \ SKIP" | Assign: @@ -183,7 +183,7 @@ computation by an antimonotonicity rule. We introduce the standard system now and show the equivalence with our formulation.\ -inductive sec_type' :: "nat \ com \ bool" ("(_/ \'' _)" [0,0] 50) where +inductive sec_type' :: "nat \ com \ bool" (\(_/ \'' _)\ [0,0] 50) where Skip': "l \' SKIP" | Assign': @@ -217,7 +217,7 @@ subsubsection "A Bottom-Up Typing System" -inductive sec_type2 :: "com \ level \ bool" ("(\ _ : _)" [0,0] 50) where +inductive sec_type2 :: "com \ level \ bool" (\(\ _ : _)\ [0,0] 50) where Skip2: "\ SKIP : l" | Assign2: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Sec_TypingT.thy --- a/src/HOL/IMP/Sec_TypingT.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Sec_TypingT.thy Fri Sep 20 19:51:08 2024 +0200 @@ -5,7 +5,7 @@ subsubsection "A Syntax Directed System" -inductive sec_type :: "nat \ com \ bool" ("(_/ \ _)" [0,0] 50) where +inductive sec_type :: "nat \ com \ bool" (\(_/ \ _)\ [0,0] 50) where Skip: "l \ SKIP" | Assign: @@ -172,7 +172,7 @@ computation by an antimonotonicity rule. We introduce the standard system now and show the equivalence with our formulation.\ -inductive sec_type' :: "nat \ com \ bool" ("(_/ \'' _)" [0,0] 50) where +inductive sec_type' :: "nat \ com \ bool" (\(_/ \'' _)\ [0,0] 50) where Skip': "l \' SKIP" | Assign': diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Sem_Equiv.thy --- a/src/HOL/IMP/Sem_Equiv.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Sem_Equiv.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,12 +9,12 @@ type_synonym assn = "state \ bool" definition - equiv_up_to :: "assn \ com \ com \ bool" ("_ \ _ \ _" [50,0,10] 50) + equiv_up_to :: "assn \ com \ com \ bool" (\_ \ _ \ _\ [50,0,10] 50) where "(P \ c \ c') = (\s s'. P s \ (c,s) \ s' \ (c',s) \ s')" definition - bequiv_up_to :: "assn \ bexp \ bexp \ bool" ("_ \ _ <\> _" [50,0,10] 50) + bequiv_up_to :: "assn \ bexp \ bexp \ bool" (\_ \ _ <\> _\ [50,0,10] 50) where "(P \ b <\> b') = (\s. P s \ bval b s = bval b' s)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Small_Step.thy --- a/src/HOL/IMP/Small_Step.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Small_Step.thy Fri Sep 20 19:51:08 2024 +0200 @@ -5,7 +5,7 @@ subsection "The transition relation" inductive - small_step :: "com * state \ com * state \ bool" (infix "\" 55) + small_step :: "com * state \ com * state \ bool" (infix \\\ 55) where Assign: "(x ::= a, s) \ (SKIP, s(x := aval a s))" | @@ -20,7 +20,7 @@ abbreviation - small_steps :: "com * state \ com * state \ bool" (infix "\*" 55) + small_steps :: "com * state \ com * state \ bool" (infix \\*\ 55) where "x \* y == star small_step x y" subsection\Executability\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Types.thy --- a/src/HOL/IMP/Types.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Types.thy Fri Sep 20 19:51:08 2024 +0200 @@ -46,16 +46,16 @@ datatype com = SKIP - | Assign vname aexp ("_ ::= _" [1000, 61] 61) - | Seq com com ("_;; _" [60, 61] 60) - | If bexp com com ("IF _ THEN _ ELSE _" [0, 0, 61] 61) - | While bexp com ("WHILE _ DO _" [0, 61] 61) + | Assign vname aexp (\_ ::= _\ [1000, 61] 61) + | Seq com com (\_;; _\ [60, 61] 60) + | If bexp com com (\IF _ THEN _ ELSE _\ [0, 0, 61] 61) + | While bexp com (\WHILE _ DO _\ [0, 61] 61) subsection "Small-Step Semantics of Commands" inductive - small_step :: "(com \ state) \ (com \ state) \ bool" (infix "\" 55) + small_step :: "(com \ state) \ (com \ state) \ bool" (infix \\\ 55) where Assign: "taval a s v \ (x ::= a, s) \ (SKIP, s(x := v))" | @@ -76,7 +76,7 @@ type_synonym tyenv = "vname \ ty" inductive atyping :: "tyenv \ aexp \ ty \ bool" - ("(1_/ \/ (_ :/ _))" [50,0,50] 50) + (\(1_/ \/ (_ :/ _))\ [50,0,50] 50) where Ic_ty: "\ \ Ic i : Ity" | Rc_ty: "\ \ Rc r : Rty" | @@ -92,7 +92,7 @@ In most situations Isabelle's type system will reject all but one parse tree, but will still inform you of the potential ambiguity.\ -inductive btyping :: "tyenv \ bexp \ bool" (infix "\" 50) +inductive btyping :: "tyenv \ bexp \ bool" (infix \\\ 50) where B_ty: "\ \ Bc v" | Not_ty: "\ \ b \ \ \ Not b" | @@ -102,7 +102,7 @@ declare btyping.intros [intro!] inductive_cases [elim!]: "\ \ Not b" "\ \ And b1 b2" "\ \ Less a1 a2" -inductive ctyping :: "tyenv \ com \ bool" (infix "\" 50) where +inductive ctyping :: "tyenv \ com \ bool" (infix \\\ 50) where Skip_ty: "\ \ SKIP" | Assign_ty: "\ \ a : \(x) \ \ \ x ::= a" | Seq_ty: "\ \ c1 \ \ \ c2 \ \ \ c1;;c2" | @@ -127,7 +127,7 @@ lemma type_eq_Rty[simp]: "type v = Rty \ (\r. v = Rv r)" by (cases v) simp_all -definition styping :: "tyenv \ state \ bool" (infix "\" 50) +definition styping :: "tyenv \ state \ bool" (infix \\\ 50) where "\ \ s \ (\x. type (s x) = \ x)" lemma apreservation: @@ -204,7 +204,7 @@ "(c,s) \ (c',s') \ \ \ c \ \ \ c'" by (induct rule: small_step_induct) (auto simp: ctyping.intros) -abbreviation small_steps :: "com * state \ com * state \ bool" (infix "\*" 55) +abbreviation small_steps :: "com * state \ com * state \ bool" (infix \\*\ 55) where "x \* y == star small_step x y" theorem type_sound: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/VCG.thy --- a/src/HOL/IMP/VCG.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/VCG.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,13 +9,13 @@ text\Commands where loops are annotated with invariants.\ datatype acom = - Askip ("SKIP") | - Aassign vname aexp ("(_ ::= _)" [1000, 61] 61) | - Aseq acom acom ("_;;/ _" [60, 61] 60) | - Aif bexp acom acom ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) | - Awhile assn bexp acom ("({_}/ WHILE _/ DO _)" [0, 0, 61] 61) + Askip (\SKIP\) | + Aassign vname aexp (\(_ ::= _)\ [1000, 61] 61) | + Aseq acom acom (\_;;/ _\ [60, 61] 60) | + Aif bexp acom acom (\(IF _/ THEN _/ ELSE _)\ [0, 0, 61] 61) | + Awhile assn bexp acom (\({_}/ WHILE _/ DO _)\ [0, 0, 61] 61) -notation com.SKIP ("SKIP") +notation com.SKIP (\SKIP\) text\Strip annotations:\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/VCG_Total_EX.thy --- a/src/HOL/IMP/VCG_Total_EX.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/VCG_Total_EX.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,14 +12,14 @@ invariants.\ datatype acom = - Askip ("SKIP") | - Aassign vname aexp ("(_ ::= _)" [1000, 61] 61) | - Aseq acom acom ("_;;/ _" [60, 61] 60) | - Aif bexp acom acom ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) | + Askip (\SKIP\) | + Aassign vname aexp (\(_ ::= _)\ [1000, 61] 61) | + Aseq acom acom (\_;;/ _\ [60, 61] 60) | + Aif bexp acom acom (\(IF _/ THEN _/ ELSE _)\ [0, 0, 61] 61) | Awhile "nat \ assn" bexp acom - ("({_}/ WHILE _/ DO _)" [0, 0, 61] 61) + (\({_}/ WHILE _/ DO _)\ [0, 0, 61] 61) -notation com.SKIP ("SKIP") +notation com.SKIP (\SKIP\) text\Strip annotations:\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/VCG_Total_EX2.thy --- a/src/HOL/IMP/VCG_Total_EX2.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/VCG_Total_EX2.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,14 +16,14 @@ invariants.\ datatype acom = - Askip ("SKIP") | - Aassign vname aexp ("(_ ::= _)" [1000, 61] 61) | - Aseq acom acom ("_;;/ _" [60, 61] 60) | - Aif bexp acom acom ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) | + Askip (\SKIP\) | + Aassign vname aexp (\(_ ::= _)\ [1000, 61] 61) | + Aseq acom acom (\_;;/ _\ [60, 61] 60) | + Aif bexp acom acom (\(IF _/ THEN _/ ELSE _)\ [0, 0, 61] 61) | Awhile assn2 lvname bexp acom - ("({_'/_}/ WHILE _/ DO _)" [0, 0, 0, 61] 61) + (\({_'/_}/ WHILE _/ DO _)\ [0, 0, 0, 61] 61) -notation com.SKIP ("SKIP") +notation com.SKIP (\SKIP\) text\Strip annotations:\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMP/Vars.thy --- a/src/HOL/IMP/Vars.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMP/Vars.thy Fri Sep 20 19:51:08 2024 +0200 @@ -49,7 +49,7 @@ abbreviation eq_on :: "('a \ 'b) \ ('a \ 'b) \ 'a set \ bool" - ("(_ =/ _/ on _)" [50,0,50] 50) where + (\(_ =/ _/ on _)\ [50,0,50] 50) where "f = g on X == \ x \ X. f x = g x" lemma aval_eq_if_eq_on_vars[simp]: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMPP/Com.thy --- a/src/HOL/IMPP/Com.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMPP/Com.thy Fri Sep 20 19:51:08 2024 +0200 @@ -33,13 +33,13 @@ datatype com = SKIP - | Ass vname aexp ("_:==_" [65, 65 ] 60) - | Local loc aexp com ("LOCAL _:=_ IN _" [65, 0, 61] 60) - | Semi com com ("_;; _" [59, 60 ] 59) - | Cond bexp com com ("IF _ THEN _ ELSE _" [65, 60, 61] 60) - | While bexp com ("WHILE _ DO _" [65, 61] 60) + | Ass vname aexp (\_:==_\ [65, 65 ] 60) + | Local loc aexp com (\LOCAL _:=_ IN _\ [65, 0, 61] 60) + | Semi com com (\_;; _\ [59, 60 ] 59) + | Cond bexp com com (\IF _ THEN _ ELSE _\ [65, 60, 61] 60) + | While bexp com (\WHILE _ DO _\ [65, 61] 60) | BODY pname - | Call vname pname aexp ("_:=CALL _'(_')" [65, 65, 0] 60) + | Call vname pname aexp (\_:=CALL _'(_')\ [65, 65, 0] 60) consts bodies :: "(pname * com) list"(* finitely many procedure definitions *) definition diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMPP/EvenOdd.thy --- a/src/HOL/IMPP/EvenOdd.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMPP/EvenOdd.thy Fri Sep 20 19:51:08 2024 +0200 @@ -35,7 +35,7 @@ end definition - Z_eq_Arg_plus :: "nat => nat assn" ("Z=Arg+_" [50]50) where + Z_eq_Arg_plus :: "nat => nat assn" (\Z=Arg+_\ [50]50) where "Z=Arg+n = (%Z s. Z = s+n)" definition diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMPP/Hoare.thy --- a/src/HOL/IMPP/Hoare.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMPP/Hoare.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,35 +25,35 @@ "state_not_singleton = (\s t::state. s ~= t)" (* at least two elements *) definition - peek_and :: "'a assn => (state => bool) => 'a assn" (infixr "&>" 35) where + peek_and :: "'a assn => (state => bool) => 'a assn" (infixr \&>\ 35) where "peek_and P p = (%Z s. P Z s & p s)" datatype 'a triple = - triple "'a assn" com "'a assn" ("{(1_)}./ (_)/ .{(1_)}" [3,60,3] 58) + triple "'a assn" com "'a assn" (\{(1_)}./ (_)/ .{(1_)}\ [3,60,3] 58) definition - triple_valid :: "nat => 'a triple => bool" ( "|=_:_" [0 , 58] 57) where + triple_valid :: "nat => 'a triple => bool" ( \|=_:_\ [0 , 58] 57) where "|=n:t = (case t of {P}.c.{Q} => \Z s. P Z s \ (\s'. -n-> s' \ Q Z s'))" abbreviation - triples_valid :: "nat => 'a triple set => bool" ("||=_:_" [0 , 58] 57) where + triples_valid :: "nat => 'a triple set => bool" (\||=_:_\ [0 , 58] 57) where "||=n:G == Ball G (triple_valid n)" definition - hoare_valids :: "'a triple set => 'a triple set => bool" ("_||=_" [58, 58] 57) where + hoare_valids :: "'a triple set => 'a triple set => bool" (\_||=_\ [58, 58] 57) where "G||=ts = (\n. ||=n:G --> ||=n:ts)" abbreviation - hoare_valid :: "'a triple set => 'a triple => bool" ("_|=_" [58, 58] 57) where + hoare_valid :: "'a triple set => 'a triple => bool" (\_|=_\ [58, 58] 57) where "G |=t == G||={t}" (* Most General Triples *) definition - MGT :: "com => state triple" ("{=}._.{->}" [60] 58) where + MGT :: "com => state triple" (\{=}._.{->}\ [60] 58) where "{=}.c.{->} = {%Z s0. Z = s0}. c .{%Z s1. -c-> s1}" inductive - hoare_derivs :: "'a triple set => 'a triple set => bool" ("_||-_" [58, 58] 57) and - hoare_deriv :: "'a triple set => 'a triple => bool" ("_|-_" [58, 58] 57) + hoare_derivs :: "'a triple set => 'a triple set => bool" (\_||-_\ [58, 58] 57) and + hoare_deriv :: "'a triple set => 'a triple => bool" (\_|-_\ [58, 58] 57) where "G |-t == G||-{t}" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IMPP/Natural.thy --- a/src/HOL/IMPP/Natural.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IMPP/Natural.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,14 +14,14 @@ newlocs :: locals setlocs :: "state => locals => state" getlocs :: "state => locals" - update :: "state => vname => val => state" ("_/[_/::=/_]" [900,0,0] 900) + update :: "state => vname => val => state" (\_/[_/::=/_]\ [900,0,0] 900) abbreviation - loc :: "state => locals" ("_<_>" [75,0] 75) where + loc :: "state => locals" (\_<_>\ [75,0] 75) where "s == getlocs s X" inductive - evalc :: "[com,state, state] => bool" ("<_,_>/ -c-> _" [0,0, 51] 51) + evalc :: "[com,state, state] => bool" (\<_,_>/ -c-> _\ [0,0, 51] 51) where Skip: " -c-> s" @@ -52,7 +52,7 @@ [X::=s1]" inductive - evaln :: "[com,state,nat,state] => bool" ("<_,_>/ -_-> _" [0,0,0,51] 51) + evaln :: "[com,state,nat,state] => bool" (\<_,_>/ -_-> _\ [0,0,0,51] 51) where Skip: " -n-> s" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/IOA/IOA.thy --- a/src/HOL/IOA/IOA.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/IOA/IOA.thy Fri Sep 20 19:51:08 2024 +0200 @@ -85,7 +85,7 @@ (outputs(a1) \ outputs(a2)), (internals(a1) \ internals(a2))))" -definition par :: "[('a,'s)ioa, ('a,'t)ioa] \ ('a,'s*'t)ioa" (infixr "||" 10) +definition par :: "[('a,'s)ioa, ('a,'t)ioa] \ ('a,'s*'t)ioa" (infixr \||\ 10) where "(ioa1 || ioa2) \ (asig_comp (asig_of ioa1) (asig_of ioa2), {pr. fst(pr) \ starts_of(ioa1) \ snd(pr) \ starts_of(ioa2)}, diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Imperative_HOL/Array.thy --- a/src/HOL/Imperative_HOL/Array.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Imperative_HOL/Array.thy Fri Sep 20 19:51:08 2024 +0200 @@ -32,7 +32,7 @@ definition update :: "'a::heap array \ nat \ 'a \ heap \ heap" where "update a i x h = set a ((get h a)[i:=x]) h" -definition noteq :: "'a::heap array \ 'b::heap array \ bool" (infix "=!!=" 70) where +definition noteq :: "'a::heap array \ 'b::heap array \ bool" (infix \=!!=\ 70) where "r =!!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_array r \ addr_of_array s" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Imperative_HOL/Overview.thy --- a/src/HOL/Imperative_HOL/Overview.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Imperative_HOL/Overview.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,12 +9,12 @@ (* type constraints with spacing *) no_syntax (output) - "_constrain" :: "logic => type => logic" ("_::_" [4, 0] 3) - "_constrain" :: "prop' => type => prop'" ("_::_" [4, 0] 3) + "_constrain" :: "logic => type => logic" (\_::_\ [4, 0] 3) + "_constrain" :: "prop' => type => prop'" (\_::_\ [4, 0] 3) syntax (output) - "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3) - "_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3) + "_constrain" :: "logic => type => logic" (\_ :: _\ [4, 0] 3) + "_constrain" :: "prop' => type => prop'" (\_ :: _\ [4, 0] 3) (*>*) text \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Imperative_HOL/Ref.thy --- a/src/HOL/Imperative_HOL/Ref.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Imperative_HOL/Ref.thy Fri Sep 20 19:51:08 2024 +0200 @@ -32,7 +32,7 @@ r = Ref l in (r, set r x (h\lim := l + 1\)))" -definition noteq :: "'a::heap ref \ 'b::heap ref \ bool" (infix "=!=" 70) where +definition noteq :: "'a::heap ref \ 'b::heap ref \ bool" (infix \=!=\ 70) where "r =!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_ref r \ addr_of_ref s" @@ -41,10 +41,10 @@ definition ref :: "'a::heap \ 'a ref Heap" where [code del]: "ref v = Heap_Monad.heap (alloc v)" -definition lookup :: "'a::heap ref \ 'a Heap" ("!_" 61) where +definition lookup :: "'a::heap ref \ 'a Heap" (\!_\ 61) where [code del]: "lookup r = Heap_Monad.tap (\h. get h r)" -definition update :: "'a ref \ 'a::heap \ unit Heap" ("_ := _" 62) where +definition update :: "'a ref \ 'a::heap \ unit Heap" (\_ := _\ 62) where [code del]: "update r v = Heap_Monad.heap (\h. ((), set r v h))" definition change :: "('a::heap \ 'a) \ 'a ref \ 'a Heap" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Induct/Com.thy --- a/src/HOL/Induct/Com.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Induct/Com.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,13 +16,13 @@ exp = N nat | X loc | Op "nat => nat => nat" exp exp - | valOf com exp ("VALOF _ RESULTIS _" 60) + | valOf com exp (\VALOF _ RESULTIS _\ 60) and com = SKIP - | Assign loc exp (infixl ":=" 60) - | Semi com com ("_;;_" [60, 60] 60) - | Cond exp com com ("IF _ THEN _ ELSE _" 60) - | While exp com ("WHILE _ DO _" 60) + | Assign loc exp (infixl \:=\ 60) + | Semi com com (\_;;_\ [60, 60] 60) + | Cond exp com com (\IF _ THEN _ ELSE _\ 60) + | While exp com (\WHILE _ DO _\ 60) subsection \Commands\ @@ -30,7 +30,7 @@ text\Execution of commands\ abbreviation (input) - generic_rel ("_/ -|[_]-> _" [50,0,50] 50) where + generic_rel (\_/ -|[_]-> _\ [50,0,50] 50) where "esig -|[eval]-> ns == (esig,ns) \ eval" text\Command execution. Natural numbers represent Booleans: 0=True, 1=False\ @@ -38,7 +38,7 @@ inductive_set exec :: "((exp*state) * (nat*state)) set => ((com*state)*state)set" and exec_rel :: "com * state => ((exp*state) * (nat*state)) set => state => bool" - ("_/ -[_]-> _" [50,0,50] 50) + (\_/ -[_]-> _\ [50,0,50] 50) for eval :: "((exp*state) * (nat*state)) set" where "csig -[eval]-> s == (csig,s) \ exec eval" @@ -108,7 +108,7 @@ inductive_set eval :: "((exp*state) * (nat*state)) set" - and eval_rel :: "[exp*state,nat*state] => bool" (infixl "-|->" 50) + and eval_rel :: "[exp*state,nat*state] => bool" (infixl \-|->\ 50) where "esig -|-> ns == (esig, ns) \ eval" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Induct/Comb.thy --- a/src/HOL/Induct/Comb.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Induct/Comb.thy Fri Sep 20 19:51:08 2024 +0200 @@ -20,14 +20,14 @@ datatype comb = K | S - | Ap comb comb (infixl "\" 90) + | Ap comb comb (infixl \\\ 90) text \ Inductive definition of contractions, \\\<^sup>1\ and (multi-step) reductions, \\\. \ -inductive contract1 :: "[comb,comb] \ bool" (infixl "\\<^sup>1" 50) +inductive contract1 :: "[comb,comb] \ bool" (infixl \\\<^sup>1\ 50) where K: "K\x\y \\<^sup>1 x" | S: "S\x\y\z \\<^sup>1 (x\z)\(y\z)" @@ -35,7 +35,7 @@ | Ap2: "x \\<^sup>1 y \ z\x \\<^sup>1 z\y" abbreviation - contract :: "[comb,comb] \ bool" (infixl "\" 50) where + contract :: "[comb,comb] \ bool" (infixl \\\ 50) where "contract \ contract1\<^sup>*\<^sup>*" text \ @@ -43,7 +43,7 @@ (multi-step) parallel reductions, \\\. \ -inductive parcontract1 :: "[comb,comb] \ bool" (infixl "\\<^sup>1" 50) +inductive parcontract1 :: "[comb,comb] \ bool" (infixl \\\<^sup>1\ 50) where refl: "x \\<^sup>1 x" | K: "K\x\y \\<^sup>1 x" @@ -51,7 +51,7 @@ | Ap: "\x \\<^sup>1 y; z \\<^sup>1 w\ \ x\z \\<^sup>1 y\w" abbreviation - parcontract :: "[comb,comb] \ bool" (infixl "\" 50) where + parcontract :: "[comb,comb] \ bool" (infixl \\\ 50) where "parcontract \ parcontract1\<^sup>*\<^sup>*" text \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Induct/Ordinals.thy --- a/src/HOL/Induct/Ordinals.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Induct/Ordinals.thy Fri Sep 20 19:51:08 2024 +0200 @@ -31,7 +31,7 @@ definition OpLim :: "(nat \ (ordinal \ ordinal)) \ (ordinal \ ordinal)" where "OpLim F a = Limit (\n. F n a)" -definition OpItw :: "(ordinal \ ordinal) \ (ordinal \ ordinal)" ("\") +definition OpItw :: "(ordinal \ ordinal) \ (ordinal \ ordinal)" (\\\) where "\f = OpLim (iter f)" primrec cantor :: "ordinal \ ordinal \ ordinal" @@ -40,7 +40,7 @@ | "cantor a (Succ b) = \(\x. cantor x b) a" | "cantor a (Limit f) = Limit (\n. cantor a (f n))" -primrec Nabla :: "(ordinal \ ordinal) \ (ordinal \ ordinal)" ("\") +primrec Nabla :: "(ordinal \ ordinal) \ (ordinal \ ordinal)" (\\\) where "\f Zero = f Zero" | "\f (Succ a) = f (Succ (\f a))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Induct/PropLog.thy --- a/src/HOL/Induct/PropLog.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Induct/PropLog.thy Fri Sep 20 19:51:08 2024 +0200 @@ -22,13 +22,13 @@ datatype 'a pl = false - | var 'a ("#_" [1000]) - | imp "'a pl" "'a pl" (infixr "\" 90) + | var 'a (\#_\ [1000]) + | imp "'a pl" "'a pl" (infixr \\\ 90) subsection \The proof system\ -inductive thms :: "['a pl set, 'a pl] \ bool" (infixl "\" 50) +inductive thms :: "['a pl set, 'a pl] \ bool" (infixl \\\ 50) for H :: "'a pl set" where H: "p \ H \ H \ p" @@ -42,7 +42,7 @@ subsubsection \Semantics of propositional logic.\ -primrec eval :: "['a set, 'a pl] => bool" ("_[[_]]" [100,0] 100) +primrec eval :: "['a set, 'a pl] => bool" (\_[[_]]\ [100,0] 100) where "tt[[false]] = False" | "tt[[#v]] = (v \ tt)" @@ -67,7 +67,7 @@ is \p\. \ -definition sat :: "['a pl set, 'a pl] => bool" (infixl "\" 50) +definition sat :: "['a pl set, 'a pl] => bool" (infixl \\\ 50) where "H \ p = (\tt. (\q\H. tt[[q]]) \ tt[[p]])" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Induct/QuoDataType.thy --- a/src/HOL/Induct/QuoDataType.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Induct/QuoDataType.thy Fri Sep 20 19:51:08 2024 +0200 @@ -30,7 +30,7 @@ inductive_set msgrel :: "(freemsg * freemsg) set" - and msg_rel :: "[freemsg, freemsg] => bool" (infixl "\" 50) + and msg_rel :: "[freemsg, freemsg] => bool" (infixl \\\ 50) where "X \ Y == (X,Y) \ msgrel" | CD: "CRYPT K (DECRYPT K X) \ X" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Induct/QuoNestedDataType.thy --- a/src/HOL/Induct/QuoNestedDataType.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Induct/QuoNestedDataType.thy Fri Sep 20 19:51:08 2024 +0200 @@ -29,7 +29,7 @@ and transitivity.\ inductive_set exprel :: "(freeExp * freeExp) set" - and exp_rel :: "[freeExp, freeExp] => bool" (infixl "\" 50) + and exp_rel :: "[freeExp, freeExp] => bool" (infixl \\\ 50) where "X \ Y \ (X,Y) \ exprel" | ASSOC: "PLUS X (PLUS Y Z) \ PLUS (PLUS X Y) Z" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Induct/SList.thy --- a/src/HOL/Induct/SList.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Induct/SList.thy Fri Sep 20 19:51:08 2024 +0200 @@ -84,15 +84,15 @@ "[x, xs]" == "x#[xs]" "[x]" == "x#[]" no_notation - Nil ("[]") and - Cons (infixr "#" 65) + Nil (\[]\) and + Cons (infixr \#\ 65) definition - Nil :: "'a list" ("[]") where + Nil :: "'a list" (\[]\) where "Nil = Abs_List(NIL)" definition - "Cons" :: "['a, 'a list] => 'a list" (infixr "#" 65) where + "Cons" :: "['a, 'a list] => 'a list" (infixr \#\ 65) where "x#xs = Abs_List(CONS (Leaf x)(Rep_List xs))" definition diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Isar_Examples/Group_Context.thy --- a/src/HOL/Isar_Examples/Group_Context.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Isar_Examples/Group_Context.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ text \hypothetical group axiomatization\ context - fixes prod :: "'a \ 'a \ 'a" (infixl "\" 70) + fixes prod :: "'a \ 'a \ 'a" (infixl \\\ 70) and one :: "'a" and inverse :: "'a \ 'a" assumes assoc: "(x \ y) \ z = x \ (y \ z)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Isar_Examples/Group_Notepad.thy --- a/src/HOL/Isar_Examples/Group_Notepad.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Isar_Examples/Group_Notepad.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ begin txt \hypothetical group axiomatization\ - fix prod :: "'a \ 'a \ 'a" (infixl "\" 70) + fix prod :: "'a \ 'a \ 'a" (infixl \\\ 70) and one :: "'a" and inverse :: "'a \ 'a" assume assoc: "(x \ y) \ z = x \ (y \ z)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Isar_Examples/Hoare.thy --- a/src/HOL/Isar_Examples/Hoare.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Isar_Examples/Hoare.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,11 +25,11 @@ datatype 'a com = Basic "'a \ 'a" - | Seq "'a com" "'a com" ("(_;/ _)" [60, 61] 60) + | Seq "'a com" "'a com" (\(_;/ _)\ [60, 61] 60) | Cond "'a bexp" "'a com" "'a com" | While "'a bexp" "'a assn" "'a var" "'a com" -abbreviation Skip ("SKIP") +abbreviation Skip (\SKIP\) where "SKIP \ Basic id" type_synonym 'a sem = "'a \ 'a \ bool" @@ -46,7 +46,7 @@ | "Sem (Cond b c1 c2) s s' \ (if s \ b then Sem c1 s s' else Sem c2 s s')" | "Sem (While b x y c) s s' \ (\n. iter n b (Sem c) s s')" -definition Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" ("(3\ _/ (2_)/ _)" [100, 55, 100] 50) +definition Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" (\(3\ _/ (2_)/ _)\ [100, 55, 100] 50) where "\ P c Q \ (\s s'. Sem c s s' \ s \ P \ s' \ Q)" lemma ValidI [intro?]: "(\s s'. Sem c s s' \ s \ P \ s' \ Q) \ \ P c Q" @@ -193,15 +193,15 @@ syntax "_quote" :: "'b \ ('a \ 'b)" - "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) - "_Subst" :: "'a bexp \ 'b \ idt \ 'a bexp" ("_[_'/\_]" [1000] 999) - "_Assert" :: "'a \ 'a set" ("(\_\)" [0] 1000) - "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) + "_antiquote" :: "('a \ 'b) \ 'b" (\\_\ [1000] 1000) + "_Subst" :: "'a bexp \ 'b \ idt \ 'a bexp" (\_[_'/\_]\ [1000] 999) + "_Assert" :: "'a \ 'a set" (\(\_\)\ [0] 1000) + "_Assign" :: "idt \ 'b \ 'a com" (\(\_ :=/ _)\ [70, 65] 61) "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" - ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61) + (\(0IF _/ THEN _/ ELSE _/ FI)\ [0, 0, 0] 61) "_While_inv" :: "'a bexp \ 'a assn \ 'a com \ 'a com" - ("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61) - "_While" :: "'a bexp \ 'a com \ 'a com" ("(0WHILE _ //DO _ /OD)" [0, 0] 61) + (\(0WHILE _/ INV _ //DO _ /OD)\ [0, 0, 0] 61) + "_While" :: "'a bexp \ 'a com \ 'a com" (\(0WHILE _ //DO _ /OD)\ [0, 0] 61) translations "\b\" \ "CONST Collect (_quote b)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Lattice/CompleteLattice.thy --- a/src/HOL/Lattice/CompleteLattice.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Lattice/CompleteLattice.thy Fri Sep 20 19:51:08 2024 +0200 @@ -31,10 +31,10 @@ \ definition - Meet :: "'a::complete_lattice set \ 'a" ("\_" [90] 90) where + Meet :: "'a::complete_lattice set \ 'a" (\\_\ [90] 90) where "\A = (THE inf. is_Inf A inf)" definition - Join :: "'a::complete_lattice set \ 'a" ("\_" [90] 90) where + Join :: "'a::complete_lattice set \ 'a" (\\_\ [90] 90) where "\A = (THE sup. is_Sup A sup)" text \ @@ -184,11 +184,11 @@ \ definition - bottom :: "'a::complete_lattice" ("\") where + bottom :: "'a::complete_lattice" (\\\) where "\ = \UNIV" definition - top :: "'a::complete_lattice" ("\") where + top :: "'a::complete_lattice" (\\\) where "\ = \UNIV" lemma bottom_least [intro?]: "\ \ x" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Lattice/Lattice.thy --- a/src/HOL/Lattice/Lattice.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Lattice/Lattice.thy Fri Sep 20 19:51:08 2024 +0200 @@ -24,10 +24,10 @@ \ definition - meet :: "'a::lattice \ 'a \ 'a" (infixl "\" 70) where + meet :: "'a::lattice \ 'a \ 'a" (infixl \\\ 70) where "x \ y = (THE inf. is_inf x y inf)" definition - join :: "'a::lattice \ 'a \ 'a" (infixl "\" 65) where + join :: "'a::lattice \ 'a \ 'a" (infixl \\\ 65) where "x \ y = (THE sup. is_sup x y sup)" text \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Lattice/Orders.thy --- a/src/HOL/Lattice/Orders.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Lattice/Orders.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,7 +17,7 @@ \ class leq = - fixes leq :: "'a \ 'a \ bool" (infixl "\" 50) + fixes leq :: "'a \ 'a \ bool" (infixl \\\ 50) class quasi_order = leq + assumes leq_refl [intro?]: "x \ x" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Cardinality.thy --- a/src/HOL/Library/Cardinality.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Cardinality.thy Fri Sep 20 19:51:08 2024 +0200 @@ -30,7 +30,7 @@ subsection \Cardinalities of types\ -syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))") +syntax "_type_card" :: "type => nat" (\(1CARD/(1'(_')))\) syntax_consts "_type_card" == card diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Combine_PER.thy --- a/src/HOL/Library/Combine_PER.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Combine_PER.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ where "combine_per P R = (\x y. P x \ P y) \ R" lemma combine_per_simp [simp]: - "combine_per P R x y \ P x \ P y \ x \ y" for R (infixl "\" 50) + "combine_per P R x y \ P x \ P y \ x \ y" for R (infixl \\\ 50) by (simp add: combine_per_def) lemma combine_per_top [simp]: "combine_per \ R = R" @@ -27,7 +27,7 @@ lemma transp_combine_per: "transp R \ transp (combine_per P R)" by (auto simp add: transp_def trans_def combine_per_def) -lemma combine_perI: "P x \ P y \ x \ y \ combine_per P R x y" for R (infixl "\" 50) +lemma combine_perI: "P x \ P y \ x \ y \ combine_per P R x y" for R (infixl \\\ 50) by (simp add: combine_per_def) lemma symp_combine_per_symp: "symp R \ symp (combine_per P R)" @@ -37,7 +37,7 @@ by (auto intro!: transpI elim: transpE) lemma equivp_combine_per_part_equivp [intro?]: - fixes R (infixl "\" 50) + fixes R (infixl \\\ 50) assumes "\x. P x" and "equivp R" shows "part_equivp (combine_per P R)" proof - diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Complete_Partial_Order2.thy --- a/src/HOL/Library/Complete_Partial_Order2.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Complete_Partial_Order2.thy Fri Sep 20 19:51:08 2024 +0200 @@ -72,7 +72,7 @@ qed lemma mono_lub: - fixes le_b (infix "\" 60) + fixes le_b (infix \\\ 60) assumes chain: "Complete_Partial_Order.chain (fun_ord (\)) Y" and mono: "\f. f \ Y \ monotone le_b (\) f" shows "monotone (\) (\) (fun_lub Sup Y)" @@ -96,7 +96,7 @@ qed context - fixes le_b (infix "\" 60) and Y f + fixes le_b (infix \\\ 60) and Y f assumes chain: "Complete_Partial_Order.chain le_b Y" and mono1: "\y. y \ Y \ monotone le_b (\) (\x. f x y)" and mono2: "\x a b. \ x \ Y; a \ b; a \ Y; b \ Y \ \ f x a \ f x b" @@ -163,7 +163,7 @@ end lemma Sup_image_mono_le: - fixes le_b (infix "\" 60) and Sup_b ("\") + fixes le_b (infix \\\ 60) and Sup_b (\\\) assumes ccpo: "class.ccpo Sup_b (\) lt_b" assumes chain: "Complete_Partial_Order.chain (\) Y" and mono: "\x y. \ x \ y; x \ Y \ \ f x \ f y" @@ -181,7 +181,7 @@ qed lemma swap_Sup: - fixes le_b (infix "\" 60) + fixes le_b (infix \\\ 60) assumes Y: "Complete_Partial_Order.chain (\) Y" and Z: "Complete_Partial_Order.chain (fun_ord (\)) Z" and mono: "\f. f \ Z \ monotone (\) (\) f" @@ -289,7 +289,7 @@ qed(blast intro: ccpo_Sup_least) qed(rule chain_iterates[OF f]) -context fixes ordb :: "'b \ 'b \ bool" (infix "\" 60) begin +context fixes ordb :: "'b \ 'b \ bool" (infix \\\ 60) begin lemma iterates_mono: assumes f: "f \ ccpo.iterates (fun_lub Sup) (fun_ord (\)) F" @@ -569,8 +569,8 @@ by(rule mcont2mcont'[OF _ mcont_const]) context - fixes ord :: "'b \ 'b \ bool" (infix "\" 60) - and lub :: "'b set \ 'b" ("\") + fixes ord :: "'b \ 'b \ bool" (infix \\\ 60) + and lub :: "'b set \ 'b" (\\\) begin lemma cont_fun_lub_Sup: @@ -691,7 +691,7 @@ by(rule monotoneI)(auto intro: bot intro: mono order_trans) lemma (in ccpo) mcont_if_bot: - fixes bot and lub ("\") and ord (infix "\" 60) + fixes bot and lub (\\\) and ord (infix \\\ 60) assumes ccpo: "class.ccpo lub (\) lt" and mono: "\x y. \ x \ y; \ x \ bound \ \ f x \ f y" and cont: "\Y. \ Complete_Partial_Order.chain (\) Y; Y \ {}; \x. x \ Y \ \ x \ bound \ \ f (\Y) = \(f ` Y)" @@ -891,7 +891,7 @@ end lemma admissible_leI: - fixes ord (infix "\" 60) and lub ("\") + fixes ord (infix \\\ 60) and lub (\\\) assumes "class.ccpo lub (\) (mk_less (\))" and "mcont luba orda lub (\) (\x. f x)" and "mcont luba orda lub (\) (\x. g x)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Datatype_Records.thy --- a/src/HOL/Library/Datatype_Records.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Datatype_Records.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,32 +27,32 @@ \ no_syntax - "_constify" :: "id => ident" ("_") - "_constify" :: "longid => ident" ("_") + "_constify" :: "id => ident" (\_\) + "_constify" :: "longid => ident" (\_\) - "_field_type" :: "ident => type => field_type" ("(2_ ::/ _)") - "" :: "field_type => field_types" ("_") - "_field_types" :: "field_type => field_types => field_types" ("_,/ _") - "_record_type" :: "field_types => type" ("(3\_\)") - "_record_type_scheme" :: "field_types => type => type" ("(3\_,/ (2\ ::/ _)\)") + "_field_type" :: "ident => type => field_type" (\(2_ ::/ _)\) + "" :: "field_type => field_types" (\_\) + "_field_types" :: "field_type => field_types => field_types" (\_,/ _\) + "_record_type" :: "field_types => type" (\(3\_\)\) + "_record_type_scheme" :: "field_types => type => type" (\(3\_,/ (2\ ::/ _)\)\) - "_field" :: "ident => 'a => field" ("(2_ =/ _)") - "" :: "field => fields" ("_") - "_fields" :: "field => fields => fields" ("_,/ _") - "_record" :: "fields => 'a" ("(3\_\)") - "_record_scheme" :: "fields => 'a => 'a" ("(3\_,/ (2\ =/ _)\)") + "_field" :: "ident => 'a => field" (\(2_ =/ _)\) + "" :: "field => fields" (\_\) + "_fields" :: "field => fields => fields" (\_,/ _\) + "_record" :: "fields => 'a" (\(3\_\)\) + "_record_scheme" :: "fields => 'a => 'a" (\(3\_,/ (2\ =/ _)\)\) - "_field_update" :: "ident => 'a => field_update" ("(2_ :=/ _)") - "" :: "field_update => field_updates" ("_") - "_field_updates" :: "field_update => field_updates => field_updates" ("_,/ _") - "_record_update" :: "'a => field_updates => 'b" ("_/(3\_\)" [900, 0] 900) + "_field_update" :: "ident => 'a => field_update" (\(2_ :=/ _)\) + "" :: "field_update => field_updates" (\_\) + "_field_updates" :: "field_update => field_updates => field_updates" (\_,/ _\) + "_record_update" :: "'a => field_updates => 'b" (\_/(3\_\)\ [900, 0] 900) no_syntax (ASCII) - "_record_type" :: "field_types => type" ("(3'(| _ |'))") - "_record_type_scheme" :: "field_types => type => type" ("(3'(| _,/ (2... ::/ _) |'))") - "_record" :: "fields => 'a" ("(3'(| _ |'))") - "_record_scheme" :: "fields => 'a => 'a" ("(3'(| _,/ (2... =/ _) |'))") - "_record_update" :: "'a => field_updates => 'b" ("_/(3'(| _ |'))" [900, 0] 900) + "_record_type" :: "field_types => type" (\(3'(| _ |'))\) + "_record_type_scheme" :: "field_types => type => type" (\(3'(| _,/ (2... ::/ _) |'))\) + "_record" :: "fields => 'a" (\(3'(| _ |'))\) + "_record_scheme" :: "fields => 'a => 'a" (\(3'(| _,/ (2... =/ _) |'))\) + "_record_update" :: "'a => field_updates => 'b" (\_/(3'(| _ |'))\ [900, 0] 900) (* copied and adapted from Record.thy *) @@ -63,22 +63,22 @@ field_updates syntax - "_constify" :: "id => ident" ("_") - "_constify" :: "longid => ident" ("_") + "_constify" :: "id => ident" (\_\) + "_constify" :: "longid => ident" (\_\) - "_datatype_field" :: "ident => 'a => field" ("(2_ =/ _)") - "" :: "field => fields" ("_") - "_datatype_fields" :: "field => fields => fields" ("_,/ _") - "_datatype_record" :: "fields => 'a" ("(3\_\)") - "_datatype_field_update" :: "ident => 'a => field_update" ("(2_ :=/ _)") - "" :: "field_update => field_updates" ("_") - "_datatype_field_updates" :: "field_update => field_updates => field_updates" ("_,/ _") - "_datatype_record_update" :: "'a => field_updates => 'b" ("_/(3\_\)" [900, 0] 900) + "_datatype_field" :: "ident => 'a => field" (\(2_ =/ _)\) + "" :: "field => fields" (\_\) + "_datatype_fields" :: "field => fields => fields" (\_,/ _\) + "_datatype_record" :: "fields => 'a" (\(3\_\)\) + "_datatype_field_update" :: "ident => 'a => field_update" (\(2_ :=/ _)\) + "" :: "field_update => field_updates" (\_\) + "_datatype_field_updates" :: "field_update => field_updates => field_updates" (\_,/ _\) + "_datatype_record_update" :: "'a => field_updates => 'b" (\_/(3\_\)\ [900, 0] 900) syntax (ASCII) - "_datatype_record" :: "fields => 'a" ("(3'(| _ |'))") - "_datatype_record_scheme" :: "fields => 'a => 'a" ("(3'(| _,/ (2... =/ _) |'))") - "_datatype_record_update" :: "'a => field_updates => 'b" ("_/(3'(| _ |'))" [900, 0] 900) + "_datatype_record" :: "fields => 'a" (\(3'(| _ |'))\) + "_datatype_record_scheme" :: "fields => 'a => 'a" (\(3'(| _,/ (2... =/ _) |'))\) + "_datatype_record_update" :: "'a => field_updates => 'b" (\_/(3'(| _ |'))\ [900, 0] 900) named_theorems datatype_record_update diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Equipollence.thy --- a/src/HOL/Library/Equipollence.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Equipollence.thy Fri Sep 20 19:51:08 2024 +0200 @@ -6,10 +6,10 @@ subsection\Eqpoll\ -definition eqpoll :: "'a set \ 'b set \ bool" (infixl "\" 50) +definition eqpoll :: "'a set \ 'b set \ bool" (infixl \\\ 50) where "eqpoll A B \ \f. bij_betw f A B" -definition lepoll :: "'a set \ 'b set \ bool" (infixl "\" 50) +definition lepoll :: "'a set \ 'b set \ bool" (infixl \\\ 50) where "lepoll A B \ \f. inj_on f A \ f ` A \ B" definition lesspoll :: "'a set \ 'b set \ bool" (infixl \\\ 50) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Extended.thy --- a/src/HOL/Library/Extended.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Extended.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,7 +8,7 @@ imports Simps_Case_Conv begin -datatype 'a extended = Fin 'a | Pinf ("\") | Minf ("-\") +datatype 'a extended = Fin 'a | Pinf (\\\) | Minf (\-\\) instantiation extended :: (order)order diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Extended_Nat.thy --- a/src/HOL/Library/Extended_Nat.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Extended_Nat.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,7 +10,7 @@ begin class infinity = - fixes infinity :: "'a" ("\") + fixes infinity :: "'a" (\\\) context fixes f :: "nat \ 'a::{canonically_ordered_monoid_add, linorder_topology, complete_linorder}" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/FSet.thy --- a/src/HOL/Library/FSet.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/FSet.thy Fri Sep 20 19:51:08 2024 +0200 @@ -58,12 +58,12 @@ end -abbreviation fempty :: "'a fset" ("{||}") where "{||} \ bot" -abbreviation fsubset_eq :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) where "xs |\| ys \ xs \ ys" -abbreviation fsubset :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) where "xs |\| ys \ xs < ys" -abbreviation funion :: "'a fset \ 'a fset \ 'a fset" (infixl "|\|" 65) where "xs |\| ys \ sup xs ys" -abbreviation finter :: "'a fset \ 'a fset \ 'a fset" (infixl "|\|" 65) where "xs |\| ys \ inf xs ys" -abbreviation fminus :: "'a fset \ 'a fset \ 'a fset" (infixl "|-|" 65) where "xs |-| ys \ minus xs ys" +abbreviation fempty :: "'a fset" (\{||}\) where "{||} \ bot" +abbreviation fsubset_eq :: "'a fset \ 'a fset \ bool" (infix \|\|\ 50) where "xs |\| ys \ xs \ ys" +abbreviation fsubset :: "'a fset \ 'a fset \ bool" (infix \|\|\ 50) where "xs |\| ys \ xs < ys" +abbreviation funion :: "'a fset \ 'a fset \ 'a fset" (infixl \|\|\ 65) where "xs |\| ys \ sup xs ys" +abbreviation finter :: "'a fset \ 'a fset \ 'a fset" (infixl \|\|\ 65) where "xs |\| ys \ inf xs ys" +abbreviation fminus :: "'a fset \ 'a fset \ 'a fset" (infixl \|-|\ 65) where "xs |-| ys \ minus xs ys" instantiation fset :: (equal) equal begin @@ -154,7 +154,7 @@ end abbreviation fUNIV :: "'a::finite fset" where "fUNIV \ top" -abbreviation fuminus :: "'a::finite fset \ 'a fset" ("|-| _" [81] 80) where "|-| x \ uminus x" +abbreviation fuminus :: "'a::finite fset \ 'a fset" (\|-| _\ [81] 80) where "|-| x \ uminus x" declare top_fset.rep_eq[simp] @@ -166,19 +166,19 @@ nonterminal fset_args syntax - "" :: "'a \ fset_args" ("_") - "_fset_args" :: "'a \ fset_args \ fset_args" ("_,/ _") - "_fset" :: "fset_args => 'a fset" ("{|(_)|}") + "" :: "'a \ fset_args" (\_\) + "_fset_args" :: "'a \ fset_args \ fset_args" (\_,/ _\) + "_fset" :: "fset_args => 'a fset" (\{|(_)|}\) syntax_consts "_fset_args" "_fset" == finsert translations "{|x, xs|}" == "CONST finsert x {|xs|}" "{|x|}" == "CONST finsert x {||}" -abbreviation fmember :: "'a \ 'a fset \ bool" (infix "|\|" 50) where +abbreviation fmember :: "'a \ 'a fset \ bool" (infix \|\|\ 50) where "x |\| X \ x \ fset X" -abbreviation not_fmember :: "'a \ 'a fset \ bool" (infix "|\|" 50) where +abbreviation not_fmember :: "'a \ 'a fset \ bool" (infix \|\|\ 50) where "x |\| X \ x \ fset X" context @@ -197,12 +197,12 @@ end syntax (input) - "_fBall" :: "pttrn \ 'a fset \ bool \ bool" ("(3! (_/|:|_)./ _)" [0, 0, 10] 10) - "_fBex" :: "pttrn \ 'a fset \ bool \ bool" ("(3? (_/|:|_)./ _)" [0, 0, 10] 10) + "_fBall" :: "pttrn \ 'a fset \ bool \ bool" (\(3! (_/|:|_)./ _)\ [0, 0, 10] 10) + "_fBex" :: "pttrn \ 'a fset \ bool \ bool" (\(3? (_/|:|_)./ _)\ [0, 0, 10] 10) syntax - "_fBall" :: "pttrn \ 'a fset \ bool \ bool" ("(3\(_/|\|_)./ _)" [0, 0, 10] 10) - "_fBex" :: "pttrn \ 'a fset \ bool \ bool" ("(3\(_/|\|_)./ _)" [0, 0, 10] 10) + "_fBall" :: "pttrn \ 'a fset \ bool \ bool" (\(3\(_/|\|_)./ _)\ [0, 0, 10] 10) + "_fBex" :: "pttrn \ 'a fset \ bool \ bool" (\(3\(_/|\|_)./ _)\ [0, 0, 10] 10) syntax_consts "_fBall" \ FSet.Ball and @@ -243,7 +243,7 @@ lift_definition fcard :: "'a fset \ nat" is card parametric card_transfer . -lift_definition fimage :: "('a \ 'b) \ 'a fset \ 'b fset" (infixr "|`|" 90) is image +lift_definition fimage :: "('a \ 'b) \ 'a fset \ 'b fset" (infixr \|`|\ 90) is image parametric image_transfer by simp lift_definition fthe_elem :: "'a fset \ 'a" is the_elem . @@ -1905,7 +1905,7 @@ text \Setup adapted from sets.\ -notation Quickcheck_Exhaustive.orelse (infixr "orelse" 55) +notation Quickcheck_Exhaustive.orelse (infixr \orelse\ 55) context includes term_syntax @@ -1939,7 +1939,7 @@ end -no_notation Quickcheck_Exhaustive.orelse (infixr "orelse" 55) +no_notation Quickcheck_Exhaustive.orelse (infixr \orelse\ 55) instantiation fset :: (random) random begin diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Finite_Map.thy --- a/src/HOL/Library/Finite_Map.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Finite_Map.thy Fri Sep 20 19:51:08 2024 +0200 @@ -497,7 +497,7 @@ lemma fmdrop_fset_fmdrop[simp]: "fmdrop_fset S (fmdrop b m) = fmdrop_fset (finsert b S) m" by (rule fmap_ext) auto -lift_definition fmadd :: "('a, 'b) fmap \ ('a, 'b) fmap \ ('a, 'b) fmap" (infixl "++\<^sub>f" 100) +lift_definition fmadd :: "('a, 'b) fmap \ ('a, 'b) fmap \ ('a, 'b) fmap" (infixl \++\<^sub>f\ 100) is map_add parametric map_add_transfer by simp @@ -611,7 +611,7 @@ obtains (none) "fmlookup m x = None" | (some) y where "fmlookup m x = Some y" "P x y" using assms by auto -lift_definition fmsubset :: "('a, 'b) fmap \ ('a, 'b) fmap \ bool" (infix "\\<^sub>f" 50) +lift_definition fmsubset :: "('a, 'b) fmap \ ('a, 'b) fmap \ bool" (infix \\\<^sub>f\ 50) is map_le . @@ -801,7 +801,7 @@ obtains x where "fmlookup m x = Some y" "x |\| A" using assms by (auto simp: fmlookup_image_iff) -lift_definition fmcomp :: "('b, 'c) fmap \ ('a, 'b) fmap \ ('a, 'c) fmap" (infixl "\\<^sub>f" 55) +lift_definition fmcomp :: "('b, 'c) fmap \ ('a, 'b) fmap \ ('a, 'c) fmap" (infixl \\\<^sub>f\ 55) is map_comp parametric map_comp_transfer by (rule dom_comp_finite) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/FuncSet.thy --- a/src/HOL/Library/FuncSet.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/FuncSet.thy Fri Sep 20 19:51:08 2024 +0200 @@ -19,12 +19,12 @@ definition "restrict" :: "('a \ 'b) \ 'a set \ 'a \ 'b" where "restrict f A = (\x. if x \ A then f x else undefined)" -abbreviation funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 60) +abbreviation funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr \\\ 60) where "A \ B \ Pi A (\_. B)" syntax - "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\ _\_./ _)" 10) - "_lam" :: "pttrn \ 'a set \ ('a \ 'b) \ ('a \ 'b)" ("(3\_\_./ _)" [0,0,3] 3) + "_Pi" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" (\(3\ _\_./ _)\ 10) + "_lam" :: "pttrn \ 'a set \ ('a \ 'b) \ ('a \ 'b)" (\(3\_\_./ _)\ [0,0,3] 3) syntax_consts "_Pi" \ Pi and "_lam" \ restrict @@ -345,13 +345,13 @@ abbreviation "Pi\<^sub>E A B \ PiE A B" syntax - "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3\\<^sub>E _\_./ _)" 10) + "_PiE" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" (\(3\\<^sub>E _\_./ _)\ 10) syntax_consts "_PiE" \ Pi\<^sub>E translations "\\<^sub>E x\A. B" \ "CONST Pi\<^sub>E A (\x. B)" -abbreviation extensional_funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\\<^sub>E" 60) +abbreviation extensional_funcset :: "'a set \ 'b set \ ('a \ 'b) set" (infixr \\\<^sub>E\ 60) where "A \\<^sub>E B \ (\\<^sub>E i\A. B)" lemma extensional_funcset_def: "extensional_funcset S T = (S \ T) \ extensional S" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Groups_Big_Fun.thy --- a/src/HOL/Library/Groups_Big_Fun.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Groups_Big_Fun.thy Fri Sep 20 19:51:08 2024 +0200 @@ -199,9 +199,9 @@ end syntax (ASCII) - "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3SUM _. _)" [0, 10] 10) + "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" (\(3SUM _. _)\ [0, 10] 10) syntax - "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3\_. _)" [0, 10] 10) + "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" (\(3\_. _)\ [0, 10] 10) syntax_consts "_Sum_any" \ Sum_any translations @@ -254,9 +254,9 @@ end syntax (ASCII) - "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3PROD _. _)" [0, 10] 10) + "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" (\(3PROD _. _)\ [0, 10] 10) syntax - "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3\_. _)" [0, 10] 10) + "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" (\(3\_. _)\ [0, 10] 10) syntax_consts "_Prod_any" == Prod_any translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/IArray.thy --- a/src/HOL/Library/IArray.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/IArray.thy Fri Sep 20 19:51:08 2024 +0200 @@ -28,7 +28,7 @@ qualified definition of_fun :: "(nat \ 'a) \ nat \ 'a iarray" where [simp]: "of_fun f n = IArray (map f [0.. nat \ 'a" (infixl "!!" 100) where +qualified definition sub :: "'a iarray \ nat \ 'a" (infixl \!!\ 100) where [simp]: "as !! n = IArray.list_of as ! n" qualified definition length :: "'a iarray \ nat" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Interval.thy --- a/src/HOL/Library/Interval.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Interval.thy Fri Sep 20 19:51:08 2024 +0200 @@ -204,7 +204,7 @@ subsection \Membership\ -abbreviation (in preorder) in_interval ("(_/ \\<^sub>i _)" [51, 51] 50) +abbreviation (in preorder) in_interval (\(_/ \\<^sub>i _)\ [51, 51] 50) where "in_interval x X \ x \ set_of X" lemma in_interval_to_interval[intro!]: "a \\<^sub>i interval_of a" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Interval_Float.thy --- a/src/HOL/Library/Interval_Float.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Interval_Float.thy Fri Sep 20 19:51:08 2024 +0200 @@ -92,7 +92,7 @@ "x \\<^sub>i X" if "lower X \ x" "x \ upper X" using that by (auto simp: set_of_eq) -abbreviation in_real_interval ("(_/ \\<^sub>r _)" [51, 51] 50) where +abbreviation in_real_interval (\(_/ \\<^sub>r _)\ [51, 51] 50) where "x \\<^sub>r X \ x \\<^sub>i real_interval X" lemma in_real_intervalI: @@ -149,15 +149,15 @@ by (auto simp: lower_def upper_def Interval_inverse split_beta') definition all_in_i :: "'a::preorder list \ 'a interval list \ bool" - (infix "(all'_in\<^sub>i)" 50) + (infix \(all'_in\<^sub>i)\ 50) where "x all_in\<^sub>i I = (length x = length I \ (\i < length I. x ! i \\<^sub>i I ! i))" definition all_in :: "real list \ float interval list \ bool" - (infix "(all'_in)" 50) + (infix \(all'_in)\ 50) where "x all_in I = (length x = length I \ (\i < length I. x ! i \\<^sub>r I ! i))" definition all_subset :: "'a::order interval list \ 'a interval list \ bool" - (infix "(all'_subset)" 50) + (infix \(all'_subset)\ 50) where "I all_subset J = (length I = length J \ (\i < length I. set_of (I!i) \ set_of (J!i)))" lemmas [simp] = all_in_def all_subset_def diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/LaTeXsugar.thy --- a/src/HOL/Library/LaTeXsugar.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/LaTeXsugar.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,28 +16,28 @@ definition mbox0 :: "'a \ 'a" where "mbox0 x = x" -notation (latex) mbox ("\<^latex>\\\mbox{\_\<^latex>\}\" [999] 999) +notation (latex) mbox (\\<^latex>\\mbox{\_\<^latex>\}\\ [999] 999) -notation (latex) mbox0 ("\<^latex>\\\mbox{\_\<^latex>\}\" [0] 999) +notation (latex) mbox0 (\\<^latex>\\mbox{\_\<^latex>\}\\ [0] 999) (* LOGIC *) notation (latex output) - If ("(\<^latex>\\\textsf{\if\<^latex>\}\ (_)/ \<^latex>\\\textsf{\then\<^latex>\}\ (_)/ \<^latex>\\\textsf{\else\<^latex>\}\ (_))" 10) + If (\(\<^latex>\\textsf{\if\<^latex>\}\ (_)/ \<^latex>\\textsf{\then\<^latex>\}\ (_)/ \<^latex>\\textsf{\else\<^latex>\}\ (_))\ 10) syntax (latex output) "_Let" :: "[letbinds, 'a] => 'a" - ("(\<^latex>\\\textsf{\let\<^latex>\}\ (_)/ \<^latex>\\\textsf{\in\<^latex>\}\ (_))" 10) + (\(\<^latex>\\textsf{\let\<^latex>\}\ (_)/ \<^latex>\\textsf{\in\<^latex>\}\ (_))\ 10) "_case_syntax":: "['a, cases_syn] => 'b" - ("(\<^latex>\\\textsf{\case\<^latex>\}\ _ \<^latex>\\\textsf{\of\<^latex>\}\/ _)" 10) + (\(\<^latex>\\textsf{\case\<^latex>\}\ _ \<^latex>\\textsf{\of\<^latex>\}\/ _)\ 10) (* SETS *) (* empty set *) notation (latex) - "Set.empty" ("\") + "Set.empty" (\\\) (* insert *) translations @@ -48,8 +48,8 @@ (* set comprehension *) syntax (latex output) - "_Collect" :: "pttrn => bool => 'a set" ("(1{_ | _})") - "_CollectIn" :: "pttrn => 'a set => bool => 'a set" ("(1{_ \ _ | _})") + "_Collect" :: "pttrn => bool => 'a set" (\(1{_ | _})\) + "_CollectIn" :: "pttrn => 'a set => bool => 'a set" (\(1{_ \ _ | _})\) translations "_Collect p P" <= "{p. P}" "_Collect p P" <= "{p|xs. P}" @@ -57,56 +57,56 @@ (* card *) notation (latex output) - card ("|_|") + card (\|_|\) (* LISTS *) (* Cons *) notation (latex) - Cons ("_ \/ _" [66,65] 65) + Cons (\_ \/ _\ [66,65] 65) (* length *) notation (latex output) - length ("|_|") + length (\|_|\) (* nth *) notation (latex output) - nth ("_\<^latex>\\\ensuremath{_{[\\mathit{\_\<^latex>\}]}}\" [1000,0] 1000) + nth (\_\<^latex>\\ensuremath{_{[\mathit{\_\<^latex>\}]}}\\ [1000,0] 1000) (* DUMMY *) -consts DUMMY :: 'a ("\<^latex>\\\_\") +consts DUMMY :: 'a (\\<^latex>\\_\\) (* THEOREMS *) notation (Rule output) - Pure.imp ("\<^latex>\\\mbox{}\\inferrule{\\mbox{\_\<^latex>\}}\\<^latex>\{\\mbox{\_\<^latex>\}}\") + Pure.imp (\\<^latex>\\mbox{}\inferrule{\mbox{\_\<^latex>\}}\\<^latex>\{\mbox{\_\<^latex>\}}\\) syntax (Rule output) "_bigimpl" :: "asms \ prop \ prop" - ("\<^latex>\\\mbox{}\\inferrule{\_\<^latex>\}\\<^latex>\{\\mbox{\_\<^latex>\}}\") + (\\<^latex>\\mbox{}\inferrule{\_\<^latex>\}\\<^latex>\{\mbox{\_\<^latex>\}}\\) "_asms" :: "prop \ asms \ asms" - ("\<^latex>\\\mbox{\_\<^latex>\}\\\\\/ _") + (\\<^latex>\\mbox{\_\<^latex>\}\\\/ _\) - "_asm" :: "prop \ asms" ("\<^latex>\\\mbox{\_\<^latex>\}\") + "_asm" :: "prop \ asms" (\\<^latex>\\mbox{\_\<^latex>\}\\) notation (Axiom output) - "Trueprop" ("\<^latex>\\\mbox{}\\inferrule{\\mbox{}}{\\mbox{\_\<^latex>\}}\") + "Trueprop" (\\<^latex>\\mbox{}\inferrule{\mbox{}}{\mbox{\_\<^latex>\}}\\) notation (IfThen output) - Pure.imp ("\<^latex>\{\\normalsize{}\If\<^latex>\\\,}\ _/ \<^latex>\{\\normalsize \\,\then\<^latex>\\\,}\/ _.") + Pure.imp (\\<^latex>\{\normalsize{}\If\<^latex>\\,}\ _/ \<^latex>\{\normalsize \,\then\<^latex>\\,}\/ _.\) syntax (IfThen output) "_bigimpl" :: "asms \ prop \ prop" - ("\<^latex>\{\\normalsize{}\If\<^latex>\\\,}\ _ /\<^latex>\{\\normalsize \\,\then\<^latex>\\\,}\/ _.") - "_asms" :: "prop \ asms \ asms" ("\<^latex>\\\mbox{\_\<^latex>\}\ /\<^latex>\{\\normalsize \\,\and\<^latex>\\\,}\/ _") - "_asm" :: "prop \ asms" ("\<^latex>\\\mbox{\_\<^latex>\}\") + (\\<^latex>\{\normalsize{}\If\<^latex>\\,}\ _ /\<^latex>\{\normalsize \,\then\<^latex>\\,}\/ _.\) + "_asms" :: "prop \ asms \ asms" (\\<^latex>\\mbox{\_\<^latex>\}\ /\<^latex>\{\normalsize \,\and\<^latex>\\,}\/ _\) + "_asm" :: "prop \ asms" (\\<^latex>\\mbox{\_\<^latex>\}\\) notation (IfThenNoBox output) - Pure.imp ("\<^latex>\{\\normalsize{}\If\<^latex>\\\,}\ _/ \<^latex>\{\\normalsize \\,\then\<^latex>\\\,}\/ _.") + Pure.imp (\\<^latex>\{\normalsize{}\If\<^latex>\\,}\ _/ \<^latex>\{\normalsize \,\then\<^latex>\\,}\/ _.\) syntax (IfThenNoBox output) "_bigimpl" :: "asms \ prop \ prop" - ("\<^latex>\{\\normalsize{}\If\<^latex>\\\,}\ _ /\<^latex>\{\\normalsize \\,\then\<^latex>\\\,}\/ _.") - "_asms" :: "prop \ asms \ asms" ("_ /\<^latex>\{\\normalsize \\,\and\<^latex>\\\,}\/ _") - "_asm" :: "prop \ asms" ("_") + (\\<^latex>\{\normalsize{}\If\<^latex>\\,}\ _ /\<^latex>\{\normalsize \,\then\<^latex>\\,}\/ _.\) + "_asms" :: "prop \ asms \ asms" (\_ /\<^latex>\{\normalsize \,\and\<^latex>\\,}\/ _\) + "_asm" :: "prop \ asms" (\_\) setup \ Document_Output.antiquotation_pretty_source_embedded \<^binding>\const_typ\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy --- a/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Fri Sep 20 19:51:08 2024 +0200 @@ -30,13 +30,13 @@ text \Propositional connectives:\ -abbreviation (input) IMPL (infix "impl" 60) +abbreviation (input) IMPL (infix \impl\ 60) where "\ impl \ \ \ xs. \ xs \ \ xs" -abbreviation (input) OR (infix "or" 60) +abbreviation (input) OR (infix \or\ 60) where "\ or \ \ \ xs. \ xs \ \ xs" -abbreviation (input) AND (infix "aand" 60) +abbreviation (input) AND (infix \aand\ 60) where "\ aand \ \ \ xs. \ xs \ \ xs" abbreviation (input) not where "not \ \ \ xs. \ \ xs" @@ -63,7 +63,7 @@ definition "HLD s = holds (\x. x \ s)" -abbreviation HLD_nxt (infixr "\" 65) where +abbreviation HLD_nxt (infixr \\\ 65) where "s \ P \ HLD s aand nxt P" context @@ -79,7 +79,7 @@ alw: "\\ xs; alw \ (stl xs)\ \ alw \ xs" \ \weak until:\ -coinductive UNTIL (infix "until" 60) for \ \ where +coinductive UNTIL (infix \until\ 60) for \ \ where base: "\ xs \ (\ until \) xs" | step: "\\ xs; (\ until \) (stl xs)\ \ (\ until \) xs" @@ -599,7 +599,7 @@ notes [[inductive_internals]] begin -inductive suntil (infix "suntil" 60) for \ \ where +inductive suntil (infix \suntil\ 60) for \ \ where base: "\ \ \ (\ suntil \) \" | step: "\ \ \ (\ suntil \) (stl \) \ (\ suntil \) \" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/ListVector.thy --- a/src/HOL/Library/ListVector.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/ListVector.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ text\Multiplication with a scalar:\ -abbreviation scale :: "('a::times) \ 'a list \ 'a list" (infix "*\<^sub>s" 70) +abbreviation scale :: "('a::times) \ 'a list \ 'a list" (infix \*\<^sub>s\ 70) where "x *\<^sub>s xs \ map ((*) x) xs" lemma scale1[simp]: "(1::'a::monoid_mult) *\<^sub>s xs = xs" @@ -99,7 +99,7 @@ subsection "Inner product" -definition iprod :: "'a::ring list \ 'a list \ 'a" ("\_,_\") where +definition iprod :: "'a::ring list \ 'a list \ 'a" (\\_,_\\) where "\xs,ys\ = (\(x,y) \ zip xs ys. x*y)" lemma iprod_Nil[simp]: "\[],ys\ = 0" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Monad_Syntax.thy --- a/src/HOL/Library/Monad_Syntax.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Monad_Syntax.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,10 +15,10 @@ \ consts - bind :: "'a \ ('b \ 'c) \ 'd" (infixl "\" 54) + bind :: "'a \ ('b \ 'c) \ 'd" (infixl \\\ 54) notation (ASCII) - bind (infixl ">>=" 54) + bind (infixl \>>=\ 54) abbreviation (do_notation) @@ -26,25 +26,25 @@ where "bind_do \ bind" notation (output) - bind_do (infixl "\" 54) + bind_do (infixl \\\ 54) notation (ASCII output) - bind_do (infixl ">>=" 54) + bind_do (infixl \>>=\ 54) nonterminal do_binds and do_bind syntax - "_do_block" :: "do_binds \ 'a" ("do {//(2 _)//}" [12] 62) - "_do_bind" :: "[pttrn, 'a] \ do_bind" ("(2_ \/ _)" 13) - "_do_let" :: "[pttrn, 'a] \ do_bind" ("(2let _ =/ _)" [1000, 13] 13) - "_do_then" :: "'a \ do_bind" ("_" [14] 13) - "_do_final" :: "'a \ do_binds" ("_") - "_do_cons" :: "[do_bind, do_binds] \ do_binds" ("_;//_" [13, 12] 12) - "_thenM" :: "['a, 'b] \ 'c" (infixl "\" 54) + "_do_block" :: "do_binds \ 'a" (\do {//(2 _)//}\ [12] 62) + "_do_bind" :: "[pttrn, 'a] \ do_bind" (\(2_ \/ _)\ 13) + "_do_let" :: "[pttrn, 'a] \ do_bind" (\(2let _ =/ _)\ [1000, 13] 13) + "_do_then" :: "'a \ do_bind" (\_\ [14] 13) + "_do_final" :: "'a \ do_binds" (\_\) + "_do_cons" :: "[do_bind, do_binds] \ do_binds" (\_;//_\ [13, 12] 12) + "_thenM" :: "['a, 'b] \ 'c" (infixl \\\ 54) syntax (ASCII) - "_do_bind" :: "[pttrn, 'a] \ do_bind" ("(2_ <-/ _)" 13) - "_thenM" :: "['a, 'b] \ 'c" (infixl ">>" 54) + "_do_bind" :: "[pttrn, 'a] \ do_bind" (\(2_ <-/ _)\ 13) + "_thenM" :: "['a, 'b] \ 'c" (infixl \>>\ 54) syntax_consts "_do_block" "_do_cons" "_do_bind" "_do_then" \ bind and diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Multiset.thy Fri Sep 20 19:51:08 2024 +0200 @@ -93,9 +93,9 @@ nonterminal multiset_args syntax - "" :: "'a \ multiset_args" ("_") - "_multiset_args" :: "'a \ multiset_args \ multiset_args" ("_,/ _") - "_multiset" :: "multiset_args \ 'a multiset" ("{#(_)#}") + "" :: "'a \ multiset_args" (\_\) + "_multiset_args" :: "'a \ multiset_args \ multiset_args" (\_,/ _\) + "_multiset" :: "multiset_args \ 'a multiset" (\{#(_)#}\) syntax_consts "_multiset_args" "_multiset" == add_mset translations @@ -167,12 +167,12 @@ end syntax - "_MBall" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\#_./ _)" [0, 0, 10] 10) - "_MBex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\#_./ _)" [0, 0, 10] 10) + "_MBall" :: "pttrn \ 'a set \ bool \ bool" (\(3\_\#_./ _)\ [0, 0, 10] 10) + "_MBex" :: "pttrn \ 'a set \ bool \ bool" (\(3\_\#_./ _)\ [0, 0, 10] 10) syntax (ASCII) - "_MBall" :: "pttrn \ 'a set \ bool \ bool" ("(3\_:#_./ _)" [0, 0, 10] 10) - "_MBex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_:#_./ _)" [0, 0, 10] 10) + "_MBall" :: "pttrn \ 'a set \ bool \ bool" (\(3\_:#_./ _)\ [0, 0, 10] 10) + "_MBex" :: "pttrn \ 'a set \ bool \ bool" (\(3\_:#_./ _)\ [0, 0, 10] 10) syntax_consts "_MBall" \ Multiset.Ball and @@ -526,27 +526,27 @@ subsubsection \Pointwise ordering induced by count\ -definition subseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) +definition subseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix \\#\ 50) where "A \# B \ (\a. count A a \ count B a)" -definition subset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) +definition subset_mset :: "'a multiset \ 'a multiset \ bool" (infix \\#\ 50) where "A \# B \ A \# B \ A \ B" -abbreviation (input) supseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) +abbreviation (input) supseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix \\#\ 50) where "supseteq_mset A B \ B \# A" -abbreviation (input) supset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) +abbreviation (input) supset_mset :: "'a multiset \ 'a multiset \ bool" (infix \\#\ 50) where "supset_mset A B \ B \# A" notation (input) - subseteq_mset (infix "\#" 50) and - supseteq_mset (infix "\#" 50) + subseteq_mset (infix \\#\ 50) and + supseteq_mset (infix \\#\ 50) notation (ASCII) - subseteq_mset (infix "<=#" 50) and - subset_mset (infix "<#" 50) and - supseteq_mset (infix ">=#" 50) and - supset_mset (infix ">#" 50) + subseteq_mset (infix \<=#\ 50) and + subset_mset (infix \<#\ 50) and + supseteq_mset (infix \>=#\ 50) and + supset_mset (infix \>#\ 50) global_interpretation subset_mset: ordering \(\#)\ \(\#)\ by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order.trans order.antisym) @@ -1254,9 +1254,9 @@ by (rule filter_preserves_multiset) syntax (ASCII) - "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{#_ :# _./ _#})") + "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" (\(1{#_ :# _./ _#})\) syntax - "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{#_ \# _./ _#})") + "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" (\(1{#_ \# _./ _#})\) syntax_consts "_MCollect" == filter_mset translations @@ -1829,18 +1829,18 @@ image_mset_subseteq_mono subset_mset.less_le_not_le) syntax (ASCII) - "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ :# _#})") + "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" (\({#_/. _ :# _#})\) syntax - "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ \# _#})") + "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" (\({#_/. _ \# _#})\) syntax_consts "_comprehension_mset" \ image_mset translations "{#e. x \# M#}" \ "CONST image_mset (\x. e) M" syntax (ASCII) - "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ :# _./ _#})") + "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" (\({#_/ | _ :# _./ _#})\) syntax - "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ \# _./ _#})") + "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" (\({#_/ | _ \# _./ _#})\) syntax_consts "_comprehension_mset'" \ image_mset translations @@ -2684,12 +2684,12 @@ end -notation sum_mset ("\\<^sub>#") +notation sum_mset (\\\<^sub>#\) syntax (ASCII) - "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3SUM _:#_. _)" [0, 51, 10] 10) + "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" (\(3SUM _:#_. _)\ [0, 51, 10] 10) syntax - "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3\_\#_. _)" [0, 51, 10] 10) + "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" (\(3\_\#_. _)\ [0, 51, 10] 10) syntax_consts "_sum_mset_image" \ sum_mset translations @@ -2865,12 +2865,12 @@ end -notation prod_mset ("\\<^sub>#") +notation prod_mset (\\\<^sub>#\) syntax (ASCII) - "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3PROD _:#_. _)" [0, 51, 10] 10) + "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" (\(3PROD _:#_. _)\ [0, 51, 10] 10) syntax - "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3\_\#_. _)" [0, 51, 10] 10) + "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" (\(3\_\#_. _)\ [0, 51, 10] 10) syntax_consts "_prod_mset_image" \ prod_mset translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Nonpos_Ints.thy --- a/src/HOL/Library/Nonpos_Ints.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Nonpos_Ints.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,7 +14,7 @@ integers \<^term>\\\) This is useful e.g. for the Gamma function. \ -definition nonpos_Ints ("\\<^sub>\\<^sub>0") where "\\<^sub>\\<^sub>0 = {of_int n |n. n \ 0}" +definition nonpos_Ints (\\\<^sub>\\<^sub>0\) where "\\<^sub>\\<^sub>0 = {of_int n |n. n \ 0}" lemma zero_in_nonpos_Ints [simp,intro]: "0 \ \\<^sub>\\<^sub>0" unfolding nonpos_Ints_def by (auto intro!: exI[of _ "0::int"]) @@ -143,7 +143,7 @@ subsection\Non-negative reals\ -definition nonneg_Reals :: "'a::real_algebra_1 set" ("\\<^sub>\\<^sub>0") +definition nonneg_Reals :: "'a::real_algebra_1 set" (\\\<^sub>\\<^sub>0\) where "\\<^sub>\\<^sub>0 = {of_real r | r. r \ 0}" lemma nonneg_Reals_of_real_iff [simp]: "of_real r \ \\<^sub>\\<^sub>0 \ r \ 0" @@ -210,7 +210,7 @@ subsection\Non-positive reals\ -definition nonpos_Reals :: "'a::real_algebra_1 set" ("\\<^sub>\\<^sub>0") +definition nonpos_Reals :: "'a::real_algebra_1 set" (\\\<^sub>\\<^sub>0\) where "\\<^sub>\\<^sub>0 = {of_real r | r. r \ 0}" lemma nonpos_Reals_of_real_iff [simp]: "of_real r \ \\<^sub>\\<^sub>0 \ r \ 0" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Numeral_Type.thy --- a/src/HOL/Library/Numeral_Type.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Numeral_Type.thy Fri Sep 20 19:51:08 2024 +0200 @@ -523,9 +523,9 @@ subsection \Syntax\ syntax - "_NumeralType" :: "num_token => type" ("_") - "_NumeralType0" :: type ("0") - "_NumeralType1" :: type ("1") + "_NumeralType" :: "num_token => type" (\_\) + "_NumeralType0" :: type (\0\) + "_NumeralType1" :: type (\1\) syntax_types "_NumeralType0" == num0 and diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Open_State_Syntax.thy --- a/src/HOL/Library/Open_State_Syntax.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Open_State_Syntax.thy Fri Sep 20 19:51:08 2024 +0200 @@ -122,15 +122,15 @@ definition "sdo_syntax = ()" syntax - "_sdo_block" :: "sdo_binds \ 'a" ("exec {//(2 _)//}" [12] 62) - "_sdo_bind" :: "[pttrn, 'a] \ sdo_bind" ("(_ \/ _)" 13) - "_sdo_let" :: "[pttrn, 'a] \ sdo_bind" ("(2let _ =/ _)" [1000, 13] 13) - "_sdo_then" :: "'a \ sdo_bind" ("_" [14] 13) - "_sdo_final" :: "'a \ sdo_binds" ("_") - "_sdo_cons" :: "[sdo_bind, sdo_binds] \ sdo_binds" ("_;//_" [13, 12] 12) + "_sdo_block" :: "sdo_binds \ 'a" (\exec {//(2 _)//}\ [12] 62) + "_sdo_bind" :: "[pttrn, 'a] \ sdo_bind" (\(_ \/ _)\ 13) + "_sdo_let" :: "[pttrn, 'a] \ sdo_bind" (\(2let _ =/ _)\ [1000, 13] 13) + "_sdo_then" :: "'a \ sdo_bind" (\_\ [14] 13) + "_sdo_final" :: "'a \ sdo_binds" (\_\) + "_sdo_cons" :: "[sdo_bind, sdo_binds] \ sdo_binds" (\_;//_\ [13, 12] 12) syntax (ASCII) - "_sdo_bind" :: "[pttrn, 'a] \ sdo_bind" ("(_ <-/ _)" 13) + "_sdo_bind" :: "[pttrn, 'a] \ sdo_bind" (\(_ <-/ _)\ 13) syntax_consts "_sdo_block" "_sdo_cons" == sdo_syntax and diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/OptionalSugar.thy --- a/src/HOL/Library/OptionalSugar.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/OptionalSugar.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,7 +27,7 @@ (* append *) syntax (latex output) - "_appendL" :: "'a list \ 'a list \ 'a list" (infixl "\<^latex>\\\isacharat\" 65) + "_appendL" :: "'a list \ 'a list \ 'a list" (infixl \\<^latex>\\isacharat\\ 65) translations "_appendL xs ys" <= "xs @ ys" "_appendL (_appendL xs ys) zs" <= "_appendL xs (_appendL ys zs)" @@ -36,8 +36,8 @@ (* deprecated, use thm with style instead, will be removed *) (* aligning equations *) notation (tab output) - "HOL.eq" ("(_) \<^latex>\}\\putisatab\\isa{\\ \=\<^latex>\}\\putisatab\\isa{\ (_)" [50,49] 50) and - "Pure.eq" ("(_) \<^latex>\}\\putisatab\\isa{\\ \\\<^latex>\}\\putisatab\\isa{\ (_)") + "HOL.eq" (\(_) \<^latex>\}\putisatab\isa{\ \=\<^latex>\}\putisatab\isa{\ (_)\ [50,49] 50) and + "Pure.eq" (\(_) \<^latex>\}\putisatab\isa{\ \\\<^latex>\}\putisatab\isa{\ (_)\) (* Let *) translations @@ -53,21 +53,21 @@ (* type constraints with spacing *) no_syntax (output) - "_constrain" :: "logic => type => logic" ("_::_" [4, 0] 3) - "_constrain" :: "prop' => type => prop'" ("_::_" [4, 0] 3) + "_constrain" :: "logic => type => logic" (\_::_\ [4, 0] 3) + "_constrain" :: "prop' => type => prop'" (\_::_\ [4, 0] 3) syntax (output) - "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3) - "_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3) + "_constrain" :: "logic => type => logic" (\_ :: _\ [4, 0] 3) + "_constrain" :: "prop' => type => prop'" (\_ :: _\ [4, 0] 3) (* sorts as intersections *) syntax (output) - "_topsort" :: "sort" ("\" 1000) - "_sort" :: "classes => sort" ("'(_')" 1000) - "_classes" :: "id => classes => classes" ("_ \ _" 1000) - "_classes" :: "longid => classes => classes" ("_ \ _" 1000) + "_topsort" :: "sort" (\\\ 1000) + "_sort" :: "classes => sort" (\'(_')\ 1000) + "_classes" :: "id => classes => classes" (\_ \ _\ 1000) + "_classes" :: "longid => classes => classes" (\_ \ _\ 1000) end (*>*) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Pattern_Aliases.thy --- a/src/HOL/Library/Pattern_Aliases.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Pattern_Aliases.thy Fri Sep 20 19:51:08 2024 +0200 @@ -160,7 +160,7 @@ bundle pattern_aliases begin - notation as (infixr "=:" 1) + notation as (infixr \=:\ 1) declaration \K (Syntax_Phases.term_check 98 "pattern_syntax" (K (map check_pattern_syntax)))\ declaration \K (Syntax_Phases.term_uncheck 98 "pattern_syntax" (map o uncheck_pattern_syntax))\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Phantom_Type.thy --- a/src/HOL/Library/Phantom_Type.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Phantom_Type.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,7 +17,7 @@ and of_phantom_comp_phantom [simp]: "of_phantom \ phantom = id" by(simp_all add: o_def id_def) -syntax "_Phantom" :: "type \ logic" ("(1Phantom/(1'(_')))") +syntax "_Phantom" :: "type \ logic" (\(1Phantom/(1'(_')))\) syntax_consts "_Phantom" == phantom translations "Phantom('t)" => "CONST phantom :: _ \ ('t, _) phantom" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Poly_Mapping.thy --- a/src/HOL/Library/Poly_Mapping.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Poly_Mapping.thy Fri Sep 20 19:51:08 2024 +0200 @@ -89,7 +89,7 @@ context zero begin -definition "when" :: "'a \ bool \ 'a" (infixl "when" 20) +definition "when" :: "'a \ bool \ 'a" (infixl \when\ 20) where "(a when P) = (if P then a else 0)" @@ -224,7 +224,7 @@ The following type is of central importance: \ -typedef (overloaded) ('a, 'b) poly_mapping ("(_ \\<^sub>0 /_)" [1, 0] 0) = +typedef (overloaded) ('a, 'b) poly_mapping (\(_ \\<^sub>0 /_)\ [1, 0] 0) = "{f :: 'a \ 'b::zero. finite {x. f x \ 0}}" morphisms lookup Abs_poly_mapping proof - diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Preorder.thy --- a/src/HOL/Library/Preorder.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Preorder.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,8 +13,8 @@ where "equiv x y \ x \ y \ y \ x" notation - equiv ("'(\')") and - equiv ("(_/ \ _)" [51, 51] 50) + equiv (\'(\')\) and + equiv (\(_/ \ _)\ [51, 51] 50) lemma equivD1: "x \ y" if "x \ y" using that by (simp add: equiv_def) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Quotient_Syntax.thy --- a/src/HOL/Library/Quotient_Syntax.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Quotient_Syntax.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,8 +9,8 @@ begin notation - rel_conj (infixr "OOO" 75) and - map_fun (infixr "--->" 55) and - rel_fun (infixr "===>" 55) + rel_conj (infixr \OOO\ 75) and + map_fun (infixr \--->\ 55) and + rel_fun (infixr \===>\ 55) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Quotient_Type.thy --- a/src/HOL/Library/Quotient_Type.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Quotient_Type.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,7 +18,7 @@ \\ :: 'a \ 'a \ bool\.\ class eqv = - fixes eqv :: "'a \ 'a \ bool" (infixl "\" 50) + fixes eqv :: "'a \ 'a \ bool" (infixl \\\ 50) class equiv = eqv + assumes equiv_refl [intro]: "x \ x" @@ -76,7 +76,7 @@ text \Abstracted equivalence classes are the canonical representation of elements of a quotient type.\ -definition "class" :: "'a::equiv \ 'a quot" ("\_\") +definition "class" :: "'a::equiv \ 'a quot" (\\_\\) where "\a\ = Abs_quot {x. a \ x}" theorem quot_exhaust: "\a. A = \a\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/RBT_Impl.thy --- a/src/HOL/Library/RBT_Impl.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/RBT_Impl.thy Fri Sep 20 19:51:08 2024 +0200 @@ -74,10 +74,10 @@ where rbt_less_prop: "rbt_less k t \ (\x\set (keys t). x < k)" -abbreviation rbt_less_symbol (infix "|\" 50) +abbreviation rbt_less_symbol (infix \|\\ 50) where "t |\ x \ rbt_less x t" -definition rbt_greater :: "'a \ ('a, 'b) rbt \ bool" (infix "\|" 50) +definition rbt_greater :: "'a \ ('a, 'b) rbt \ bool" (infix \\|\ 50) where rbt_greater_prop: "rbt_greater k t = (\x\set (keys t). k < x)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Ramsey.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,7 +15,7 @@ subsubsection \The $n$-element subsets of a set $A$\ -definition nsets :: "['a set, nat] \ 'a set set" ("([_]\<^bsup>_\<^esup>)" [0,999] 999) +definition nsets :: "['a set, nat] \ 'a set set" (\([_]\<^bsup>_\<^esup>)\ [0,999] 999) where "nsets A n \ {N. N \ A \ finite N \ card N = n}" lemma finite_imp_finite_nsets: "finite A \ finite ([A]\<^bsup>k\<^esup>)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Real_Mod.thy --- a/src/HOL/Library/Real_Mod.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Real_Mod.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,7 +9,7 @@ (* MOVED TO HOL-Library ON 20.03.2024 *) -definition rmod :: "real \ real \ real" (infixl "rmod" 70) where +definition rmod :: "real \ real \ real" (infixl \rmod\ 70) where "x rmod y = x - \y\ * of_int \x / \y\\" lemma rmod_conv_frac: "y \ 0 \ x rmod y = frac (x / \y\) * \y\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Rewrite.thy --- a/src/HOL/Library/Rewrite.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Rewrite.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,7 +10,7 @@ imports Main begin -consts rewrite_HOLE :: "'a::{}" ("\") +consts rewrite_HOLE :: "'a::{}" (\\\) lemma eta_expand: fixes f :: "'a::{} \ 'b::{}" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Set_Algebras.thy --- a/src/HOL/Library/Set_Algebras.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Set_Algebras.thy Fri Sep 20 19:51:08 2024 +0200 @@ -54,13 +54,13 @@ end -definition elt_set_plus :: "'a::plus \ 'a set \ 'a set" (infixl "+o" 70) +definition elt_set_plus :: "'a::plus \ 'a set \ 'a set" (infixl \+o\ 70) where "a +o B = {c. \b\B. c = a + b}" -definition elt_set_times :: "'a::times \ 'a set \ 'a set" (infixl "*o" 80) +definition elt_set_times :: "'a::times \ 'a set \ 'a set" (infixl \*o\ 80) where "a *o B = {c. \b\B. c = a * b}" -abbreviation (input) elt_set_eq :: "'a \ 'a set \ bool" (infix "=o" 50) +abbreviation (input) elt_set_eq :: "'a \ 'a set \ bool" (infix \=o\ 50) where "x =o A \ x \ A" instance set :: (semigroup_add) semigroup_add diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Set_Idioms.thy --- a/src/HOL/Library/Set_Idioms.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Set_Idioms.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,11 +13,11 @@ subsection\Idioms for being a suitable union/intersection of something\ definition union_of :: "('a set set \ bool) \ ('a set \ bool) \ 'a set \ bool" - (infixr "union'_of" 60) + (infixr \union'_of\ 60) where "P union_of Q \ \S. \\. P \ \ \ \ Collect Q \ \\ = S" definition intersection_of :: "('a set set \ bool) \ ('a set \ bool) \ 'a set \ bool" - (infixr "intersection'_of" 60) + (infixr \intersection'_of\ 60) where "P intersection_of Q \ \S. \\. P \ \ \ \ Collect Q \ \\ = S" definition arbitrary:: "'a set set \ bool" where "arbitrary \ \ True" @@ -310,7 +310,7 @@ text\A somewhat cheap but handy way of getting localized forms of various topological concepts (open, closed, borel, fsigma, gdelta etc.)\ -definition relative_to :: "['a set \ bool, 'a set, 'a set] \ bool" (infixl "relative'_to" 55) +definition relative_to :: "['a set \ bool, 'a set, 'a set] \ bool" (infixl \relative'_to\ 55) where "P relative_to S \ \T. \U. P U \ S \ U = T" lemma relative_to_UNIV [simp]: "(P relative_to UNIV) S \ P S" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Sublist.thy --- a/src/HOL/Library/Sublist.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Sublist.thy Fri Sep 20 19:51:08 2024 +0200 @@ -473,7 +473,7 @@ subsection \Parallel lists\ -definition parallel :: "'a list \ 'a list \ bool" (infixl "\" 50) +definition parallel :: "'a list \ 'a list \ bool" (infixl \\\ 50) where "(xs \ ys) = (\ prefix xs ys \ \ prefix ys xs)" lemma parallelI [intro]: "\ prefix xs ys \ \ prefix ys xs \ xs \ ys" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Tree.thy --- a/src/HOL/Library/Tree.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Tree.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,8 +8,8 @@ begin datatype 'a tree = - Leaf ("\\") | - Node "'a tree" ("value": 'a) "'a tree" ("(1\_,/ _,/ _\)") + Leaf (\\\\) | + Node "'a tree" ("value": 'a) "'a tree" (\(1\_,/ _,/ _\)\) datatype_compat tree primrec left :: "'a tree \ 'a tree" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Library/Word.thy --- a/src/HOL/Library/Word.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Library/Word.thy Fri Sep 20 19:51:08 2024 +0200 @@ -1684,13 +1684,13 @@ by (simp flip: signed_take_bit_decr_length_iff) notation - word_sle ("'(\s')") and - word_sle ("(_/ \s _)" [51, 51] 50) and - word_sless ("'('(\s')\) and + word_sle (\(_/ \s _)\ [51, 51] 50) and + word_sless (\'() and + word_sless (\(_/ [51, 51] 50) notation (input) - word_sle ("(_/ <=s _)" [51, 51] 50) + word_sle (\(_/ <=s _)\ [51, 51] 50) lemma word_sle_eq [code]: \a <=s b \ sint a \ sint b\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Metis_Examples/Big_O.thy --- a/src/HOL/Metis_Examples/Big_O.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Metis_Examples/Big_O.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ subsection \Definitions\ -definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where +definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" (\(1O'(_'))\) where "O(f::('a => 'b)) == {h. \c. \x. \h x\ <= c * \f x\}" lemma bigo_pos_const: @@ -610,7 +610,7 @@ subsection \Less than or equal to\ -definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl " 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl \ 70) where "f x. max (f x - g x) 0)" lemma bigo_lesseq1: "f =o O(h) \ \x. \g x\ <= \f x\ \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Metis_Examples/Message.thy --- a/src/HOL/Metis_Examples/Message.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Metis_Examples/Message.thy Fri Sep 20 19:51:08 2024 +0200 @@ -50,9 +50,9 @@ text\Concrete syntax: messages appear as \\A,B,NA\\, etc...\ nonterminal mtuple_args syntax - "" :: "'a \ mtuple_args" ("_") - "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" ("_,/ _") - "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" ("(2\_,/ _\)") + "" :: "'a \ mtuple_args" (\_\) + "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" (\_,/ _\) + "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" (\(2\_,/ _\)\) syntax_consts "_MTuple_args" "_MTuple" \ MPair translations @@ -60,7 +60,7 @@ "\x, y\" \ "CONST MPair x y" -definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where +definition HPair :: "[msg,msg] => msg" (\(4Hash[_] /_)\ [0, 1000]) where \ \Message Y paired with a MAC computed with the help of X\ "Hash[X] Y == \ Hash\X,Y\, Y\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Metis_Examples/Tarski.thy --- a/src/HOL/Metis_Examples/Tarski.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Metis_Examples/Tarski.thy Fri Sep 20 19:51:08 2024 +0200 @@ -78,7 +78,7 @@ \pset = S, order = induced S (order cl)\ \ CompleteLattice}" abbreviation - sublattice_syntax :: "['a set, 'a potype] => bool" ("_ <<= _" [51, 50] 50) + sublattice_syntax :: "['a set, 'a potype] => bool" (\_ <<= _\ [51, 50] 50) where "S <<= cl \ S \ sublattice `` {cl}" definition dual :: "'a potype => 'a potype" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/BV/BVExample.thy --- a/src/HOL/MicroJava/BV/BVExample.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/BV/BVExample.thy Fri Sep 20 19:51:08 2024 +0200 @@ -163,7 +163,7 @@ @{prop [display] "P n"} \ -definition intervall :: "nat \ nat \ nat \ bool" ("_ \ [_, _')") where +definition intervall :: "nat \ nat \ nat \ bool" (\_ \ [_, _')\) where "x \ [a, b) \ a \ x \ x < b" lemma pc_0: "x < n \ (x \ [0, n) \ P x) \ P x" @@ -233,7 +233,7 @@ lemmas eff_simps [simp] = eff_def norm_eff_def xcpt_eff_def declare appInvoke [simp del] -definition phi_append :: method_type ("\\<^sub>a") where +definition phi_append :: method_type (\\\<^sub>a\) where "\\<^sub>a \ map (\(x,y). Some (x, map OK y)) [ ( [], [Class list_name, Class list_name]), ( [Class list_name], [Class list_name, Class list_name]), @@ -295,7 +295,7 @@ abbreviation Ctest :: ty where "Ctest == Class test_name" -definition phi_makelist :: method_type ("\\<^sub>m") where +definition phi_makelist :: method_type (\\\<^sub>m\) where "\\<^sub>m \ map (\(x,y). Some (x, y)) [ ( [], [OK Ctest, Err , Err ]), ( [Clist], [OK Ctest, Err , Err ]), @@ -369,7 +369,7 @@ done text \The whole program is welltyped:\ -definition Phi :: prog_type ("\") where +definition Phi :: prog_type (\\\) where "\ C sg \ if C = test_name \ sg = (makelist_name, []) then \\<^sub>m else if C = list_name \ sg = (append_name, [Class list_name]) then \\<^sub>a else []" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/BV/BVNoTypeError.thy --- a/src/HOL/MicroJava/BV/BVNoTypeError.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/BV/BVNoTypeError.thy Fri Sep 20 19:51:08 2024 +0200 @@ -426,7 +426,7 @@ done corollary no_type_errors_initial: - fixes G ("\") and Phi ("\") + fixes G (\\\) and Phi (\\\) assumes wt: "wt_jvm_prog \ \" assumes is_class: "is_class \ C" and "method": "method (\,C) (m,[]) = Some (C, b)" @@ -447,7 +447,7 @@ state or in the canonical start state) \ corollary welltyped_commutes: - fixes G ("\") and Phi ("\") + fixes G (\\\) and Phi (\\\) assumes wt: "wt_jvm_prog \ \" and *: "\,\ \JVM s \" shows "\ \ (Normal s) \jvmd\ (Normal t) = \ \ s \jvm\ t" apply rule @@ -458,7 +458,7 @@ done corollary welltyped_initial_commutes: - fixes G ("\") and Phi ("\") + fixes G (\\\) and Phi (\\\) assumes wt: "wt_jvm_prog \ \" assumes is_class: "is_class \ C" and "method": "method (\,C) (m,[]) = Some (C, b)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/BV/BVSpecTypeSafe.thy --- a/src/HOL/MicroJava/BV/BVSpecTypeSafe.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/BV/BVSpecTypeSafe.thy Fri Sep 20 19:51:08 2024 +0200 @@ -1319,7 +1319,7 @@ done lemma - fixes G :: jvm_prog ("\") + fixes G :: jvm_prog (\\\) assumes wf: "wf_prog wf_mb \" shows hconf_start: "\ \h (start_heap \) \" apply (unfold hconf_def start_heap_def) @@ -1330,7 +1330,7 @@ done lemma - fixes G :: jvm_prog ("\") and Phi :: prog_type ("\") + fixes G :: jvm_prog (\\\) and Phi :: prog_type (\\\) shows BV_correct_initial: "wt_jvm_prog \ \ \ is_class \ C \ method (\,C) (m,[]) = Some (C, b) \ \,\ \JVM start_state G C m \" @@ -1347,8 +1347,8 @@ done theorem typesafe: - fixes G :: jvm_prog ("\") - and Phi :: prog_type ("\") + fixes G :: jvm_prog (\\\) + and Phi :: prog_type (\\\) assumes welltyped: "wt_jvm_prog \ \" and main_method: "is_class \ C" "method (\,C) (m,[]) = Some (C, b)" and exec_all: "G \ start_state \ C m \jvm\ s" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/BV/Correct.thy --- a/src/HOL/MicroJava/BV/Correct.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/BV/Correct.thy Fri Sep 20 19:51:08 2024 +0200 @@ -42,7 +42,7 @@ correct_frames G hp phi rT sig frs))))" definition correct_state :: "[jvm_prog,prog_type,jvm_state] \ bool" - ("_,_ \JVM _ \" [51,51] 50) where + (\_,_ \JVM _ \\ [51,51] 50) where "correct_state G phi == \(xp,hp,frs). case xp of None \ (case frs of diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/BV/JVMType.thy --- a/src/HOL/MicroJava/BV/JVMType.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/BV/JVMType.thy Fri Sep 20 19:51:08 2024 +0200 @@ -38,19 +38,19 @@ "sup S maxs maxr == snd(snd(sl S maxs maxr))" definition sup_ty_opt :: "['code prog,ty err,ty err] \ bool" - ("_ \ _ <=o _" [71,71] 70) where + (\_ \ _ <=o _\ [71,71] 70) where "sup_ty_opt G == Err.le (subtype G)" definition sup_loc :: "['code prog,locvars_type,locvars_type] \ bool" - ("_ \ _ <=l _" [71,71] 70) where + (\_ \ _ <=l _\ [71,71] 70) where "sup_loc G == Listn.le (sup_ty_opt G)" definition sup_state :: "['code prog,state_type,state_type] \ bool" - ("_ \ _ <=s _" [71,71] 70) where + (\_ \ _ <=s _\ [71,71] 70) where "sup_state G == Product.le (Listn.le (subtype G)) (sup_loc G)" definition sup_state_opt :: "['code prog,state_type option,state_type option] \ bool" - ("_ \ _ <='' _" [71,71] 70) where + (\_ \ _ <='' _\ [71,71] 70) where "sup_state_opt G == Opt.le (sup_state G)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/Comp/CorrComp.thy --- a/src/HOL/MicroJava/Comp/CorrComp.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/Comp/CorrComp.thy Fri Sep 20 19:51:08 2024 +0200 @@ -85,7 +85,7 @@ bytecode \ aheap \ opstack \ locvars \ bool" - ("{_,_,_} \ {_, _, _} >- _ \ {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60) where + (\{_,_,_} \ {_, _, _} >- _ \ {_, _, _}\ [61,61,61,61,61,61,90,61,61,61]60) where "{G,C,S} \ {hp0, os0, lvars0} >- instrs \ {hp1, os1, lvars1} == \pre post frs. (gis (gmb G C S) = pre @ instrs @ post) \ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/Comp/TranslCompTp.thy --- a/src/HOL/MicroJava/Comp/TranslCompTp.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/Comp/TranslCompTp.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,7 +9,7 @@ (**********************************************************************) definition comb :: "['a \ 'b list \ 'c, 'c \ 'b list \ 'd, 'a] \ 'b list \ 'd" - (infixr "\" 55) + (infixr \\\ 55) where "comb == (\ f1 f2 x0. let (xs1, x1) = f1 x0; (xs2, x2) = f2 x1 diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/DFA/LBVComplete.thy --- a/src/HOL/MicroJava/DFA/LBVComplete.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/DFA/LBVComplete.thy Fri Sep 20 19:51:08 2024 +0200 @@ -24,7 +24,7 @@ locale lbvc = lbv + - fixes phi :: "'a list" ("\") + fixes phi :: "'a list" (\\\) fixes c :: "'a list" defines cert_def: "c \ make_cert step \ \" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/DFA/LBVCorrect.thy --- a/src/HOL/MicroJava/DFA/LBVCorrect.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/DFA/LBVCorrect.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,10 +9,10 @@ begin locale lbvs = lbv + - fixes s0 :: 'a ("s\<^sub>0") + fixes s0 :: 'a (\s\<^sub>0\) fixes c :: "'a list" fixes ins :: "'b list" - fixes phi :: "'a list" ("\") + fixes phi :: "'a list" (\\\) defines phi_def: "\ \ map (\pc. if c!pc = \ then wtl (take pc ins) c 0 s0 else c!pc) [0.. \x. B <=_r x" locale lbv = Semilat + - fixes T :: "'a" ("\") - fixes B :: "'a" ("\") + fixes T :: "'a" (\\\) + fixes B :: "'a" (\\\) fixes step :: "'a step_type" assumes top: "top r \" assumes T_A: "\ \ A" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/DFA/Listn.thy --- 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 \ 'a ord \ 'a list \ bool" - ("(_ /<=[_] _)" [50, 0, 51] 50) + (\(_ /<=[_] _)\ [50, 0, 51] 50) where "x <=[r] y == x <=_(le r) y" abbreviation lesssublist_syntax :: "'a list \ 'a ord \ 'a list \ bool" - ("(_ /<[_] _)" [50, 0, 51] 50) + (\(_ /<[_] _)\ [50, 0, 51] 50) where "x <[r] y == x <_(le r) y" definition map2 :: "('a \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" where @@ -30,7 +30,7 @@ abbreviation plussublist_syntax :: "'a list \ ('a \ 'b \ 'c) \ 'b list \ 'c list" - ("(_ /+[_] _)" [65, 0, 66] 65) + (\(_ /+[_] _)\ [65, 0, 66] 65) where "x +[f] y == x +_(map2 f) y" primrec coalesce :: "'a err list \ 'a list err" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/DFA/Product.thy --- a/src/HOL/MicroJava/DFA/Product.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/DFA/Product.thy Fri Sep 20 19:51:08 2024 +0200 @@ -20,7 +20,7 @@ abbreviation lesubprod_sntax :: "'a * 'b \ 'a ord \ 'b ord \ 'a * 'b \ bool" - ("(_ /<='(_,_') _)" [50, 0, 0, 51] 50) + (\(_ /<='(_,_') _)\ [50, 0, 0, 51] 50) where "p <=(rA,rB) q == p <=_(le rA rB) q" lemma unfold_lesub_prod: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/DFA/Semilat.thy --- a/src/HOL/MicroJava/DFA/Semilat.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/DFA/Semilat.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,26 +25,26 @@ where "plussub x f y = f x y" notation (ASCII) - "lesub" ("(_ /<='__ _)" [50, 1000, 51] 50) and - "lesssub" ("(_ /<'__ _)" [50, 1000, 51] 50) and - "plussub" ("(_ /+'__ _)" [65, 1000, 66] 65) + "lesub" (\(_ /<='__ _)\ [50, 1000, 51] 50) and + "lesssub" (\(_ /<'__ _)\ [50, 1000, 51] 50) and + "plussub" (\(_ /+'__ _)\ [65, 1000, 66] 65) notation - "lesub" ("(_ /\\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and - "lesssub" ("(_ /\\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and - "plussub" ("(_ /\\<^bsub>_\<^esub> _)" [65, 0, 66] 65) + "lesub" (\(_ /\\<^bsub>_\<^esub> _)\ [50, 0, 51] 50) and + "lesssub" (\(_ /\\<^bsub>_\<^esub> _)\ [50, 0, 51] 50) and + "plussub" (\(_ /\\<^bsub>_\<^esub> _)\ [65, 0, 66] 65) (* allow \ instead of \..\ *) abbreviation (input) - lesub1 :: "'a \ 'a ord \ 'a \ bool" ("(_ /\\<^sub>_ _)" [50, 1000, 51] 50) + lesub1 :: "'a \ 'a ord \ 'a \ bool" (\(_ /\\<^sub>_ _)\ [50, 1000, 51] 50) where "x \\<^sub>r y == x \\<^bsub>r\<^esub> y" abbreviation (input) - lesssub1 :: "'a \ 'a ord \ 'a \ bool" ("(_ /\\<^sub>_ _)" [50, 1000, 51] 50) + lesssub1 :: "'a \ 'a ord \ 'a \ bool" (\(_ /\\<^sub>_ _)\ [50, 1000, 51] 50) where "x \\<^sub>r y == x \\<^bsub>r\<^esub> y" abbreviation (input) - plussub1 :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" ("(_ /\\<^sub>_ _)" [65, 1000, 66] 65) + plussub1 :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" (\(_ /\\<^sub>_ _)\ [65, 1000, 66] 65) where "x \\<^sub>f y == x \\<^bsub>f\<^esub> y" definition ord :: "('a \ 'a) set \ 'a ord" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/DFA/SemilatAlg.thy --- a/src/HOL/MicroJava/DFA/SemilatAlg.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,10 +10,10 @@ begin definition lesubstep_type :: "(nat \ 's) list \ 's ord \ (nat \ 's) list \ bool" - ("(_ /\|_| _)" [50, 0, 51] 50) where + (\(_ /\|_| _)\ [50, 0, 51] 50) where "x \|r| y \ \(p,s) \ set x. \s'. (p,s') \ set y \ s <=_r s'" -primrec plusplussub :: "'a list \ ('a \ 'a \ 'a) \ 'a \ 'a" ("(_ /++'__ _)" [65, 1000, 66] 65) where +primrec plusplussub :: "'a list \ ('a \ 'a \ 'a) \ 'a \ 'a" (\(_ /++'__ _)\ [65, 1000, 66] 65) where "[] ++_f y = y" | "(x#xs) ++_f y = xs ++_f (x +_f y)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/J/Conform.thy --- a/src/HOL/MicroJava/J/Conform.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/J/Conform.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,27 +9,27 @@ type_synonym 'c env' = "'c prog \ (vname \ ty)" \ \same as \env\ of \WellType.thy\\ -definition hext :: "aheap => aheap => bool" ("_ \| _" [51,51] 50) where +definition hext :: "aheap => aheap => bool" (\_ \| _\ [51,51] 50) where "h\|h' == \a C fs. h a = Some(C,fs) --> (\fs'. h' a = Some(C,fs'))" definition conf :: "'c prog => aheap => val => ty => bool" - ("_,_ \ _ ::\ _" [51,51,51,51] 50) where + (\_,_ \ _ ::\ _\ [51,51,51,51] 50) where "G,h\v::\T == \T'. typeof (map_option obj_ty o h) v = Some T' \ G\T'\T" definition lconf :: "'c prog => aheap => ('a \ val) => ('a \ ty) => bool" - ("_,_ \ _ [::\] _" [51,51,51,51] 50) where + (\_,_ \ _ [::\] _\ [51,51,51,51] 50) where "G,h\vs[::\]Ts == \n T. Ts n = Some T --> (\v. vs n = Some v \ G,h\v::\T)" -definition oconf :: "'c prog => aheap => obj => bool" ("_,_ \ _ \" [51,51,51] 50) where +definition oconf :: "'c prog => aheap => obj => bool" (\_,_ \ _ \\ [51,51,51] 50) where "G,h\obj \ == G,h\snd obj[::\]map_of (fields (G,fst obj))" -definition hconf :: "'c prog => aheap => bool" ("_ \h _ \" [51,51] 50) where +definition hconf :: "'c prog => aheap => bool" (\_ \h _ \\ [51,51] 50) where "G\h h \ == \a obj. h a = Some obj --> G,h\obj \" definition xconf :: "aheap \ val option \ bool" where "xconf hp vo == preallocated hp \ (\ v. (vo = Some v) \ (\ xc. v = (Addr (XcptRef xc))))" -definition conforms :: "xstate => java_mb env' => bool" ("_ ::\ _" [51,51] 50) where +definition conforms :: "xstate => java_mb env' => bool" (\_ ::\ _\ [51,51] 50) where "s::\E == prg E\h heap (store s) \ \ prg E,heap (store s)\locals (store s)[::\]localT E \ xconf (heap (store s)) (abrupt s)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/J/Eval.thy --- a/src/HOL/MicroJava/J/Eval.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/J/Eval.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,13 +10,13 @@ \ \Auxiliary notions\ -definition fits :: "java_mb prog \ state \ val \ ty \ bool" ("_,_\_ fits _"[61,61,61,61]60) where +definition fits :: "java_mb prog \ state \ val \ ty \ bool" (\_,_\_ fits _\[61,61,61,61]60) where "G,s\a' fits T \ case T of PrimT T' \ False | RefT T' \ a'=Null \ G\obj_ty(lookup_obj s a')\T" -definition catch :: "java_mb prog \ xstate \ cname \ bool" ("_,_\catch _"[61,61,61]60) where +definition catch :: "java_mb prog \ xstate \ cname \ bool" (\_,_\catch _\[61,61,61]60) where "G,s\catch C\ case abrupt s of None \ False | Some a \ G,store s\ a fits Class C" -definition lupd :: "vname \ val \ state \ state" ("lupd'(_\_')"[10,10]1000) where +definition lupd :: "vname \ val \ state \ state" (\lupd'(_\_')\[10,10]1000) where "lupd vn v \ \ (hp,loc). (hp, (loc(vn\v)))" definition new_xcpt_var :: "vname \ xstate \ xstate" where @@ -27,12 +27,12 @@ inductive eval :: "[java_mb prog,xstate,expr,val,xstate] => bool " - ("_ \ _ -_\_-> _" [51,82,60,82,82] 81) + (\_ \ _ -_\_-> _\ [51,82,60,82,82] 81) and evals :: "[java_mb prog,xstate,expr list, val list,xstate] => bool " - ("_ \ _ -_[\]_-> _" [51,82,60,51,82] 81) + (\_ \ _ -_[\]_-> _\ [51,82,60,51,82] 81) and exec :: "[java_mb prog,xstate,stmt, xstate] => bool " - ("_ \ _ -_-> _" [51,82,60,82] 81) + (\_ \ _ -_-> _\ [51,82,60,82] 81) for G :: "java_mb prog" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/J/Term.thy --- a/src/HOL/MicroJava/J/Term.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/J/Term.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,21 +14,21 @@ | Lit val \ \literal value, also references\ | BinOp binop expr expr \ \binary operation\ | LAcc vname \ \local (incl. parameter) access\ - | LAss vname expr ("_::=_" [90,90]90) \ \local assign\ - | FAcc cname expr vname ("{_}_.._" [10,90,99]90) \ \field access\ + | LAss vname expr (\_::=_\ [90,90]90) \ \local assign\ + | FAcc cname expr vname (\{_}_.._\ [10,90,99]90) \ \field access\ | FAss cname expr vname - expr ("{_}_.._:=_" [10,90,99,90]90) \ \field ass.\ + expr (\{_}_.._:=_\ [10,90,99,90]90) \ \field ass.\ | Call cname expr mname - "ty list" "expr list" ("{_}_.._'( {_}_')" [10,90,99,10,10] 90) \ \method call\ + "ty list" "expr list" (\{_}_.._'( {_}_')\ [10,90,99,10,10] 90) \ \method call\ datatype_compat expr datatype stmt = Skip \ \empty statement\ | Expr expr \ \expression statement\ - | Comp stmt stmt ("_;; _" [61,60]60) - | Cond expr stmt stmt ("If '(_') _ Else _" [80,79,79]70) - | Loop expr stmt ("While '(_') _" [80,79]70) + | Comp stmt stmt (\_;; _\ [61,60]60) + | Cond expr stmt stmt (\If '(_') _ Else _\ [80,79,79]70) + | Loop expr stmt (\While '(_') _\ [80,79]70) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/J/TypeRel.thy --- a/src/HOL/MicroJava/J/TypeRel.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/J/TypeRel.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,14 +12,14 @@ inductive_set subcls1 :: "'c prog => (cname \ cname) set" - and subcls1' :: "'c prog => cname \ cname => bool" ("_ \ _ \C1 _" [71,71,71] 70) + and subcls1' :: "'c prog => cname \ cname => bool" (\_ \ _ \C1 _\ [71,71,71] 70) for G :: "'c prog" where "G \ C \C1 D \ (C, D) \ subcls1 G" | subcls1I: "\class G C = Some (D,rest); C \ Object\ \ G \ C \C1 D" abbreviation - subcls :: "'c prog => cname \ cname => bool" ("_ \ _ \C _" [71,71,71] 70) + subcls :: "'c prog => cname \ cname => bool" (\_ \ _ \C _\ [71,71,71] 70) where "G \ C \C D \ (C, D) \ (subcls1 G)\<^sup>*" lemma subcls1D: @@ -217,7 +217,7 @@ \ \widening, viz. method invocation conversion,cf. 5.3 i.e. sort of syntactic subtyping\ inductive - widen :: "'c prog => [ty , ty ] => bool" ("_ \ _ \ _" [71,71,71] 70) + widen :: "'c prog => [ty , ty ] => bool" (\_ \ _ \ _\ [71,71,71] 70) for G :: "'c prog" where refl [intro!, simp]: "G\ T \ T" \ \identity conv., cf. 5.1.1\ @@ -231,7 +231,7 @@ \ \casting conversion, cf. 5.5 / 5.1.5\ \ \left out casts on primitve types\ inductive - cast :: "'c prog => [ty , ty ] => bool" ("_ \ _ \? _" [71,71,71] 70) + cast :: "'c prog => [ty , ty ] => bool" (\_ \ _ \? _\ [71,71,71] 70) for G :: "'c prog" where widen: "G\ C\ D ==> G\C \? D" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/J/WellType.thy --- a/src/HOL/MicroJava/J/WellType.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/J/WellType.thy Fri Sep 20 19:51:08 2024 +0200 @@ -100,9 +100,9 @@ \ \local variables might include This, which is hidden anyway\ inductive - ty_expr :: "'c env => expr => ty => bool" ("_ \ _ :: _" [51, 51, 51] 50) - and ty_exprs :: "'c env => expr list => ty list => bool" ("_ \ _ [::] _" [51, 51, 51] 50) - and wt_stmt :: "'c env => stmt => bool" ("_ \ _ \" [51, 51] 50) + ty_expr :: "'c env => expr => ty => bool" (\_ \ _ :: _\ [51, 51, 51] 50) + and ty_exprs :: "'c env => expr list => ty list => bool" (\_ \ _ [::] _\ [51, 51, 51] 50) + and wt_stmt :: "'c env => stmt => bool" (\_ \ _ \\ [51, 51] 50) where NewC: "[| is_class (prg E) C |] ==> diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/JVM/JVMDefensive.thy --- a/src/HOL/MicroJava/JVM/JVMDefensive.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/JVM/JVMDefensive.thy Fri Sep 20 19:51:08 2024 +0200 @@ -119,7 +119,7 @@ definition exec_all_d :: "jvm_prog \ jvm_state type_error \ jvm_state type_error \ bool" - ("_ \ _ \jvmd\ _" [61,61,61]60) where + (\_ \ _ \jvmd\ _\ [61,61,61]60) where "G \ s \jvmd\ t \ (s,t) \ ({(s,t). exec_d G s = TypeError \ t = TypeError} \ {(s,t). \t'. exec_d G s = Normal (Some t') \ t = Normal t'})\<^sup>*" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/MicroJava/JVM/JVMExec.thy --- a/src/HOL/MicroJava/JVM/JVMExec.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/MicroJava/JVM/JVMExec.thy Fri Sep 20 19:51:08 2024 +0200 @@ -24,7 +24,7 @@ definition exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool" - ("_ \ _ \jvm\ _" [61,61,61]60) where + (\_ \ _ \jvm\ _\ [61,61,61]60) where "G \ s \jvm\ t == (s,t) \ {(s,t). exec(G,s) = Some t}\<^sup>*" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/NanoJava/AxSem.thy --- a/src/HOL/NanoJava/AxSem.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/NanoJava/AxSem.thy Fri Sep 20 19:51:08 2024 +0200 @@ -20,12 +20,12 @@ subsection "Hoare Logic Rules" inductive - hoare :: "[triple set, triple set] => bool" ("_ |\/ _" [61, 61] 60) - and ehoare :: "[triple set, etriple] => bool" ("_ |\\<^sub>e/ _" [61, 61] 60) + hoare :: "[triple set, triple set] => bool" (\_ |\/ _\ [61, 61] 60) + and ehoare :: "[triple set, etriple] => bool" (\_ |\\<^sub>e/ _\ [61, 61] 60) and hoare1 :: "[triple set, assn,stmt,assn] => bool" - ("_ \/ ({(1_)}/ (_)/ {(1_)})" [61, 3, 90, 3] 60) + (\_ \/ ({(1_)}/ (_)/ {(1_)})\ [61, 3, 90, 3] 60) and ehoare1 :: "[triple set, assn,expr,vassn]=> bool" - ("_ \\<^sub>e/ ({(1_)}/ (_)/ {(1_)})" [61, 3, 90, 3] 60) + (\_ \\<^sub>e/ ({(1_)}/ (_)/ {(1_)})\ [61, 3, 90, 3] 60) where "A \ {P}c{Q} \ A |\ {(P,c,Q)}" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/NanoJava/Equivalence.thy --- a/src/HOL/NanoJava/Equivalence.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/NanoJava/Equivalence.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,25 +9,25 @@ subsection "Validity" -definition valid :: "[assn,stmt, assn] => bool" ("\ {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) where +definition valid :: "[assn,stmt, assn] => bool" (\\ {(1_)}/ (_)/ {(1_)}\ [3,90,3] 60) where "\ {P} c {Q} \ \s t. P s --> (\n. s -c -n\ t) --> Q t" -definition evalid :: "[assn,expr,vassn] => bool" ("\\<^sub>e {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) where +definition evalid :: "[assn,expr,vassn] => bool" (\\\<^sub>e {(1_)}/ (_)/ {(1_)}\ [3,90,3] 60) where "\\<^sub>e {P} e {Q} \ \s v t. P s --> (\n. s -e\v-n\ t) --> Q v t" -definition nvalid :: "[nat, triple ] => bool" ("\_: _" [61,61] 60) where +definition nvalid :: "[nat, triple ] => bool" (\\_: _\ [61,61] 60) where "\n: t \ let (P,c,Q) = t in \s t. s -c -n\ t --> P s --> Q t" -definition envalid :: "[nat,etriple ] => bool" ("\_:\<^sub>e _" [61,61] 60) where +definition envalid :: "[nat,etriple ] => bool" (\\_:\<^sub>e _\ [61,61] 60) where "\n:\<^sub>e t \ let (P,e,Q) = t in \s v t. s -e\v-n\ t --> P s --> Q v t" -definition nvalids :: "[nat, triple set] => bool" ("|\_: _" [61,61] 60) where +definition nvalids :: "[nat, triple set] => bool" (\|\_: _\ [61,61] 60) where "|\n: T \ \t\T. \n: t" -definition cnvalids :: "[triple set,triple set] => bool" ("_ |\/ _" [61,61] 60) where +definition cnvalids :: "[triple set,triple set] => bool" (\_ |\/ _\ [61,61] 60) where "A |\ C \ \n. |\n: A --> |\n: C" -definition cenvalid :: "[triple set,etriple ] => bool" ("_ |\\<^sub>e/ _"[61,61] 60) where +definition cenvalid :: "[triple set,etriple ] => bool" (\_ |\\<^sub>e/ _\[61,61] 60) where "A |\\<^sub>e t \ \n. |\n: A --> \n:\<^sub>e t" lemma nvalid_def2: "\n: (P,c,Q) \ \s t. s -c-n\ t \ P s \ Q t" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/NanoJava/Example.thy --- a/src/HOL/NanoJava/Example.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/NanoJava/Example.thy Fri Sep 20 19:51:08 2024 +0200 @@ -49,13 +49,13 @@ subsection "Program representation" axiomatization - N :: cname ("Nat") (* with mixfix because of clash with NatDef.Nat *) + N :: cname (\Nat\) (* with mixfix because of clash with NatDef.Nat *) and pred :: fname and suc add :: mname and any :: vname abbreviation - dummy :: expr ("<>") + dummy :: expr (\<>\) where "<> == LAcc any" abbreviation @@ -92,7 +92,7 @@ subsection "``atleast'' relation for interpretation of Nat ``values''" -primrec Nat_atleast :: "state \ val \ nat \ bool" ("_:_ \ _" [51, 51, 51] 50) where +primrec Nat_atleast :: "state \ val \ nat \ bool" (\_:_ \ _\ [51, 51, 51] 50) where "s:x\0 = (x\Null)" | "s:x\Suc n = (\a. x=Addr a \ heap s a \ None \ s:get_field s a pred\n)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/NanoJava/OpSem.thy --- a/src/HOL/NanoJava/OpSem.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/NanoJava/OpSem.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,8 +8,8 @@ theory OpSem imports State begin inductive - exec :: "[state,stmt, nat,state] => bool" ("_ -_-_\ _" [98,90, 65,98] 89) - and eval :: "[state,expr,val,nat,state] => bool" ("_ -_\_-_\ _"[98,95,99,65,98] 89) + exec :: "[state,stmt, nat,state] => bool" (\_ -_-_\ _\ [98,90, 65,98] 89) + and eval :: "[state,expr,val,nat,state] => bool" (\_ -_\_-_\ _\[98,95,99,65,98] 89) where Skip: " s -Skip-n\ s" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/NanoJava/State.thy --- a/src/HOL/NanoJava/State.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/NanoJava/State.thy Fri Sep 20 19:51:08 2024 +0200 @@ -56,7 +56,7 @@ definition set_locs :: "state => state => state" where "set_locs s s' \ s' (| locals := locals s |)" -definition get_local :: "state => vname => val" ("_<_>" [99,0] 99) where +definition get_local :: "state => vname => val" (\_<_>\ [99,0] 99) where "get_local s x \ the (locals s x)" \ \local function:\ @@ -70,10 +70,10 @@ "get_field s a f \ the (snd (get_obj s a) f)" \ \local function:\ -definition hupd :: "loc => obj => state => state" ("hupd'(_\_')" [10,10] 1000) where +definition hupd :: "loc => obj => state => state" (\hupd'(_\_')\ [10,10] 1000) where "hupd a obj s \ s (| heap := ((heap s)(a\obj))|)" -definition lupd :: "vname => val => state => state" ("lupd'(_\_')" [10,10] 1000) where +definition lupd :: "vname => val => state => state" (\lupd'(_\_')\ [10,10] 1000) where "lupd x v s \ s (| locals := ((locals s)(x\v ))|)" definition new_obj :: "loc => cname => state => state" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/NanoJava/Term.thy --- a/src/HOL/NanoJava/Term.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/NanoJava/Term.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,20 +25,20 @@ datatype stmt = Skip \ \empty statement\ - | Comp stmt stmt ("_;; _" [91,90 ] 90) - | Cond expr stmt stmt ("If '(_') _ Else _" [ 3,91,91] 91) - | Loop vname stmt ("While '(_') _" [ 3,91 ] 91) - | LAss vname expr ("_ :== _" [99, 95] 94) \ \local assignment\ - | FAss expr fname expr ("_.._:==_" [95,99,95] 94) \ \field assignment\ + | Comp stmt stmt (\_;; _\ [91,90 ] 90) + | Cond expr stmt stmt (\If '(_') _ Else _\ [ 3,91,91] 91) + | Loop vname stmt (\While '(_') _\ [ 3,91 ] 91) + | LAss vname expr (\_ :== _\ [99, 95] 94) \ \local assignment\ + | FAss expr fname expr (\_.._:==_\ [95,99,95] 94) \ \field assignment\ | Meth "cname \ mname" \ \virtual method\ | Impl "cname \ mname" \ \method implementation\ and expr - = NewC cname ("new _" [ 99] 95) \ \object creation\ + = NewC cname (\new _\ [ 99] 95) \ \object creation\ | Cast cname expr \ \type cast\ | LAcc vname \ \local access\ - | FAcc expr fname ("_.._" [95,99] 95) \ \field access\ + | FAcc expr fname (\_.._\ [95,99] 95) \ \field access\ | Call cname expr mname expr - ("{_}_.._'(_')" [99,95,99,95] 95) \ \method call\ + (\{_}_.._'(_')\ [99,95,99,95] 95) \ \method call\ end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/NanoJava/TypeRel.thy --- a/src/HOL/NanoJava/TypeRel.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/NanoJava/TypeRel.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,10 +15,10 @@ "subcls1 \ {(C,D). C\Object \ (\c. class C = Some c \ super c=D)}" abbreviation - subcls1_syntax :: "[cname, cname] => bool" ("_ \C1 _" [71,71] 70) + subcls1_syntax :: "[cname, cname] => bool" (\_ \C1 _\ [71,71] 70) where "C \C1 D == (C,D) \ subcls1" abbreviation - subcls_syntax :: "[cname, cname] => bool" ("_ \C _" [71,71] 70) + subcls_syntax :: "[cname, cname] => bool" (\_ \C _\ [71,71] 70) where "C \C D \ (C,D) \ subcls1\<^sup>*" @@ -26,7 +26,7 @@ text\Widening, viz. method invocation conversion\ inductive - widen :: "ty => ty => bool" ("_ \ _" [71,71] 70) + widen :: "ty => ty => bool" (\_ \ _\ [71,71] 70) where refl [intro!, simp]: "T \ T" | subcls: "C\C D \ Class C \ Class D" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/CK_Machine.thy --- a/src/HOL/Nominal/Examples/CK_Machine.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/CK_Machine.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,21 +25,21 @@ nominal_datatype lam = VAR "name" | APP "lam" "lam" -| LAM "\name\lam" ("LAM [_]._") +| LAM "\name\lam" (\LAM [_]._\) | NUM "nat" -| DIFF "lam" "lam" ("_ -- _") (* subtraction *) -| PLUS "lam" "lam" ("_ ++ _") (* addition *) +| DIFF "lam" "lam" (\_ -- _\) (* subtraction *) +| PLUS "lam" "lam" (\_ ++ _\) (* addition *) | TRUE | FALSE | IF "lam" "lam" "lam" -| FIX "\name\lam" ("FIX [_]._") (* recursion *) +| FIX "\name\lam" (\FIX [_]._\) (* recursion *) | ZET "lam" (* zero test *) | EQI "lam" "lam" (* equality test on numbers *) section \Capture-Avoiding Substitution\ nominal_primrec - subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\ [100,100,100] 100) where "(VAR x)[y::=s] = (if x=y then s else (VAR x))" | "(APP t\<^sub>1 t\<^sub>2)[y::=s] = APP (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])" @@ -81,7 +81,7 @@ section \Evaluation Contexts\ datatype ctx = - Hole ("\") + Hole (\\\) | CAPPL "ctx" "lam" | CAPPR "lam" "ctx" | CDIFFL "ctx" "lam" @@ -96,7 +96,7 @@ text \The operation of filling a term into a context:\ fun - filling :: "ctx \ lam \ lam" ("_\_\") + filling :: "ctx \ lam \ lam" (\_\_\\) where "\\t\ = t" | "(CAPPL E t')\t\ = APP (E\t\) t'" @@ -113,7 +113,7 @@ text \The operation of composing two contexts:\ fun - ctx_compose :: "ctx \ ctx \ ctx" ("_ \ _") + ctx_compose :: "ctx \ ctx \ ctx" (\_ \ _\) where "\ \ E' = E'" | "(CAPPL E t') \ E' = CAPPL (E \ E') t'" @@ -134,7 +134,7 @@ text \Composing a list (stack) of contexts.\ fun - ctx_composes :: "ctx list \ ctx" ("_\") + ctx_composes :: "ctx list \ ctx" (\_\\) where "[]\ = \" | "(E#Es)\ = (Es\) \ E" @@ -152,7 +152,7 @@ equivariance val inductive - machine :: "lam\ctx list\lam\ctx list\bool" ("<_,_> \ <_,_>") + machine :: "lam\ctx list\lam\ctx list\bool" (\<_,_> \ <_,_>\) where m1[intro]: " \ e2)#Es>" | m2[intro]: "val v \ e2)#Es> \ )#Es>" @@ -176,7 +176,7 @@ | mH[intro]: "n1\n2 \ )#Es> \ " inductive - "machine_star" :: "lam\ctx list\lam\ctx list\bool" ("<_,_> \* <_,_>") + "machine_star" :: "lam\ctx list\lam\ctx list\bool" (\<_,_> \* <_,_>\) where ms1[intro]: " \* " | ms2[intro]: "\ \ ; \* \ \ \* " @@ -194,7 +194,7 @@ section \The Evaluation Relation (Big-Step Semantics)\ inductive - eval :: "lam\lam\bool" ("_ \ _") + eval :: "lam\lam\bool" (\_ \ _\) where eval_NUM[intro]: "NUM n \ NUM n" | eval_DIFF[intro]: "\t1 \ (NUM n1); t2 \ (NUM n2)\ \ t1 -- t2 \ NUM (n1 - n2)" @@ -257,7 +257,7 @@ by (simp add: perm_nat_def perm_bool) inductive - cbv :: "lam\lam\bool" ("_ \cbv _") + cbv :: "lam\lam\bool" (\_ \cbv _\) where cbv1: "\val v; x\v\ \ APP (LAM [x].t) v \cbv t[x::=v]" | cbv2[intro]: "t \cbv t' \ APP t t2 \cbv APP t' t2" @@ -297,7 +297,7 @@ qed inductive - "cbv_star" :: "lam\lam\bool" (" _ \cbv* _") + "cbv_star" :: "lam\lam\bool" (\ _ \cbv* _\) where cbvs1[intro]: "e \cbv* e" | cbvs2[intro]: "\e1\cbv e2; e2 \cbv* e3\ \ e1 \cbv* e3" @@ -366,7 +366,7 @@ tVAR "string" | tBOOL | tINT -| tARR "ty" "ty" ("_ \ _") +| tARR "ty" "ty" (\_ \ _\) declare ty.inject[simp] @@ -384,7 +384,7 @@ text \Sub-Typing Contexts\ abbreviation - "sub_tctx" :: "tctx \ tctx \ bool" ("_ \ _") + "sub_tctx" :: "tctx \ tctx \ bool" (\_ \ _\) where "\\<^sub>1 \ \\<^sub>2 \ \x. x \ set \\<^sub>1 \ x \ set \\<^sub>2" @@ -425,7 +425,7 @@ section \The Typing Relation\ inductive - typing :: "tctx \ lam \ ty \ bool" ("_ \ _ : _") + typing :: "tctx \ lam \ ty \ bool" (\_ \ _ : _\) where t_VAR[intro]: "\valid \; (x,T)\set \\ \ \ \ VAR x : T" | t_APP[intro]: "\\ \ t\<^sub>1 : T\<^sub>1\T\<^sub>2; \ \ t\<^sub>2 : T\<^sub>1\ \ \ \ APP t\<^sub>1 t\<^sub>2 : T\<^sub>2" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/CR.thy --- a/src/HOL/Nominal/Examples/CR.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/CR.thy Fri Sep 20 19:51:08 2024 +0200 @@ -146,7 +146,7 @@ section \Beta Reduction\ inductive - "Beta" :: "lam\lam\bool" (" _ \\<^sub>\ _" [80,80] 80) + "Beta" :: "lam\lam\bool" (\ _ \\<^sub>\ _\ [80,80] 80) where b1[intro]: "s1\\<^sub>\s2 \ (App s1 t)\\<^sub>\(App s2 t)" | b2[intro]: "s1\\<^sub>\s2 \ (App t s1)\\<^sub>\(App t s2)" @@ -159,7 +159,7 @@ by (simp_all add: abs_fresh fresh_fact') inductive - "Beta_star" :: "lam\lam\bool" (" _ \\<^sub>\\<^sup>* _" [80,80] 80) + "Beta_star" :: "lam\lam\bool" (\ _ \\<^sub>\\<^sup>* _\ [80,80] 80) where bs1[intro, simp]: "M \\<^sub>\\<^sup>* M" | bs2[intro]: "\M1\\<^sub>\\<^sup>* M2; M2 \\<^sub>\ M3\ \ M1 \\<^sub>\\<^sup>* M3" @@ -176,7 +176,7 @@ section \One-Reduction\ inductive - One :: "lam\lam\bool" (" _ \\<^sub>1 _" [80,80] 80) + One :: "lam\lam\bool" (\ _ \\<^sub>1 _\ [80,80] 80) where o1[intro!]: "M\\<^sub>1M" | o2[simp,intro!]: "\t1\\<^sub>1t2;s1\\<^sub>1s2\ \ (App t1 s1)\\<^sub>1(App t2 s2)" @@ -189,7 +189,7 @@ by (simp_all add: abs_fresh fresh_fact') inductive - "One_star" :: "lam\lam\bool" (" _ \\<^sub>1\<^sup>* _" [80,80] 80) + "One_star" :: "lam\lam\bool" (\ _ \\<^sub>1\<^sup>* _\ [80,80] 80) where os1[intro, simp]: "M \\<^sub>1\<^sup>* M" | os2[intro]: "\M1\\<^sub>1\<^sup>* M2; M2 \\<^sub>1 M3\ \ M1 \\<^sub>1\<^sup>* M3" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/CR_Takahashi.thy --- a/src/HOL/Nominal/Examples/CR_Takahashi.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/CR_Takahashi.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,10 +16,10 @@ nominal_datatype lam = Var "name" | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + | Lam "\name\lam" (\Lam [_]._\ [100,100] 100) nominal_primrec - subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\ [100,100,100] 100) where "(Var x)[y::=s] = (if x=y then s else (Var x))" | "(App t\<^sub>1 t\<^sub>2)[y::=s] = App (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])" @@ -66,7 +66,7 @@ section \Beta-Reduction\ inductive - "Beta" :: "lam\lam\bool" (" _ \\<^sub>\ _" [80,80] 80) + "Beta" :: "lam\lam\bool" (\ _ \\<^sub>\ _\ [80,80] 80) where b1[intro]: "t1 \\<^sub>\ t2 \ App t1 s \\<^sub>\ App t2 s" | b2[intro]: "s1 \\<^sub>\ s2 \ App t s1 \\<^sub>\ App t s2" @@ -76,7 +76,7 @@ section \Transitive Closure of Beta\ inductive - "Beta_star" :: "lam\lam\bool" (" _ \\<^sub>\\<^sup>* _" [80,80] 80) + "Beta_star" :: "lam\lam\bool" (\ _ \\<^sub>\\<^sup>* _\ [80,80] 80) where bs1[intro]: "t \\<^sub>\\<^sup>* t" | bs2[intro]: "t \\<^sub>\ s \ t \\<^sub>\\<^sup>* s" @@ -85,7 +85,7 @@ section \One-Reduction\ inductive - One :: "lam\lam\bool" (" _ \\<^sub>1 _" [80,80] 80) + One :: "lam\lam\bool" (\ _ \\<^sub>1 _\ [80,80] 80) where o1[intro]: "Var x\\<^sub>1 Var x" | o2[intro]: "\t1\\<^sub>1t2; s1\\<^sub>1s2\ \ App t1 s1 \\<^sub>1 App t2 s2" @@ -148,7 +148,7 @@ section \Transitive Closure of One\ inductive - "One_star" :: "lam\lam\bool" (" _ \\<^sub>1\<^sup>* _" [80,80] 80) + "One_star" :: "lam\lam\bool" (\ _ \\<^sub>1\<^sup>* _\ [80,80] 80) where os1[intro]: "t \\<^sub>1\<^sup>* t" | os2[intro]: "t \\<^sub>1 s \ t \\<^sub>1\<^sup>* s" @@ -157,7 +157,7 @@ section \Complete Development Reduction\ inductive - Dev :: "lam \ lam \ bool" (" _ \\<^sub>d _" [80,80]80) + Dev :: "lam \ lam \ bool" (\ _ \\<^sub>d _\ [80,80]80) where d1[intro]: "Var x \\<^sub>d Var x" | d2[intro]: "t \\<^sub>d s \ Lam [x].t \\<^sub>d Lam[x].s" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Class1.thy --- a/src/HOL/Nominal/Examples/Class1.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Class1.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,14 +8,14 @@ text \types\ -no_notation not ("NOT") +no_notation not (\NOT\) nominal_datatype ty = PR "string" | NOT "ty" - | AND "ty" "ty" ("_ AND _" [100,100] 100) - | OR "ty" "ty" ("_ OR _" [100,100] 100) - | IMP "ty" "ty" ("_ IMP _" [100,100] 100) + | AND "ty" "ty" (\_ AND _\ [100,100] 100) + | OR "ty" "ty" (\_ OR _\ [100,100] 100) + | IMP "ty" "ty" (\_ IMP _\ [100,100] 100) instantiation ty :: size begin @@ -50,30 +50,30 @@ nominal_datatype trm = Ax "name" "coname" - | Cut "\coname\trm" "\name\trm" ("Cut <_>._ ('(_'))._" [0,0,0,100] 101) - | NotR "\name\trm" "coname" ("NotR ('(_'))._ _" [0,100,100] 101) - | NotL "\coname\trm" "name" ("NotL <_>._ _" [0,100,100] 101) - | AndR "\coname\trm" "\coname\trm" "coname" ("AndR <_>._ <_>._ _" [0,100,0,100,100] 101) - | AndL1 "\name\trm" "name" ("AndL1 (_)._ _" [100,100,100] 101) - | AndL2 "\name\trm" "name" ("AndL2 (_)._ _" [100,100,100] 101) - | OrR1 "\coname\trm" "coname" ("OrR1 <_>._ _" [100,100,100] 101) - | OrR2 "\coname\trm" "coname" ("OrR2 <_>._ _" [100,100,100] 101) - | OrL "\name\trm" "\name\trm" "name" ("OrL (_)._ (_)._ _" [100,100,100,100,100] 101) - | ImpR "\name\(\coname\trm)" "coname" ("ImpR (_).<_>._ _" [100,100,100,100] 101) - | ImpL "\coname\trm" "\name\trm" "name" ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 101) + | Cut "\coname\trm" "\name\trm" (\Cut <_>._ ('(_'))._\ [0,0,0,100] 101) + | NotR "\name\trm" "coname" (\NotR ('(_'))._ _\ [0,100,100] 101) + | NotL "\coname\trm" "name" (\NotL <_>._ _\ [0,100,100] 101) + | AndR "\coname\trm" "\coname\trm" "coname" (\AndR <_>._ <_>._ _\ [0,100,0,100,100] 101) + | AndL1 "\name\trm" "name" (\AndL1 (_)._ _\ [100,100,100] 101) + | AndL2 "\name\trm" "name" (\AndL2 (_)._ _\ [100,100,100] 101) + | OrR1 "\coname\trm" "coname" (\OrR1 <_>._ _\ [100,100,100] 101) + | OrR2 "\coname\trm" "coname" (\OrR2 <_>._ _\ [100,100,100] 101) + | OrL "\name\trm" "\name\trm" "name" (\OrL (_)._ (_)._ _\ [100,100,100,100,100] 101) + | ImpR "\name\(\coname\trm)" "coname" (\ImpR (_).<_>._ _\ [100,100,100,100] 101) + | ImpL "\coname\trm" "\name\trm" "name" (\ImpL <_>._ (_)._ _\ [100,100,100,100,100] 101) text \named terms\ -nominal_datatype ntrm = Na "\name\trm" ("((_):_)" [100,100] 100) +nominal_datatype ntrm = Na "\name\trm" (\((_):_)\ [100,100] 100) text \conamed terms\ -nominal_datatype ctrm = Co "\coname\trm" ("(<_>:_)" [100,100] 100) +nominal_datatype ctrm = Co "\coname\trm" (\(<_>:_)\ [100,100] 100) text \renaming functions\ nominal_primrec (freshness_context: "(d::coname,e::coname)") - crename :: "trm \ coname \ coname \ trm" ("_[_\c>_]" [100,0,0] 100) + crename :: "trm \ coname \ coname \ trm" (\_[_\c>_]\ [100,0,0] 100) where "(Ax x a)[d\c>e] = (if a=d then Ax x e else Ax x a)" | "\a\(d,e,N);x\M\ \ (Cut .M (x).N)[d\c>e] = Cut .(M[d\c>e]) (x).(N[d\c>e])" @@ -92,7 +92,7 @@ by(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess)+ nominal_primrec (freshness_context: "(u::name,v::name)") - nrename :: "trm \ name \ name \ trm" ("_[_\n>_]" [100,0,0] 100) + nrename :: "trm \ name \ name \ trm" (\_[_\n>_]\ [100,0,0] 100) where "(Ax x a)[u\n>v] = (if x=u then Ax v a else Ax x a)" | "\a\N;x\(u,v,M)\ \ (Cut .M (x).N)[u\n>v] = Cut .(M[u\n>v]) (x).(N[u\n>v])" @@ -610,7 +610,7 @@ qed nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)") - substn :: "trm \ name \ coname \ trm \ trm" ("_{_:=<_>._}" [100,100,100,100] 100) + substn :: "trm \ name \ coname \ trm \ trm" (\_{_:=<_>._}\ [100,100,100,100] 100) where "(Ax x a){y:=.P} = (if x=y then Cut .P (y).Ax y a else Ax x a)" | "\a\(c,P,N);x\(y,P,M)\ \ (Cut .M (x).N){y:=.P} = @@ -673,7 +673,7 @@ done nominal_primrec (freshness_context: "(d::name,z::coname,P::trm)") - substc :: "trm \ coname \ name \ trm \ trm" ("_{_:=('(_'))._}" [100,0,0,100] 100) + substc :: "trm \ coname \ name \ trm \ trm" (\_{_:=('(_'))._}\ [100,0,0,100] 100) where "(Ax x a){d:=(z).P} = (if d=a then Cut .(Ax x a) (z).P else Ax x a)" | "\a\(d,P,N);x\(z,P,M)\ \ (Cut .M (x).N){d:=(z).P} = @@ -2423,7 +2423,7 @@ qed (use fic_rest_elims in force)+ inductive - l_redu :: "trm \ trm \ bool" ("_ \\<^sub>l _" [100,100] 100) + l_redu :: "trm \ trm \ bool" (\_ \\<^sub>l _\ [100,100] 100) where LAxR: "\x\M; a\b; fic M a\ \ Cut .M (x).(Ax x b) \\<^sub>l M[a\c>b]" | LAxL: "\a\M; x\y; fin M x\ \ Cut .(Ax y a) (x).M \\<^sub>l M[x\n>y]" @@ -3044,7 +3044,7 @@ inductive - c_redu :: "trm \ trm \ bool" ("_ \\<^sub>c _" [100,100] 100) + c_redu :: "trm \ trm \ bool" (\_ \\<^sub>c _\ [100,100] 100) where left[intro]: "\\fic M a; a\N; x\M\ \ Cut .M (x).N \\<^sub>c M{a:=(x).N}" | right[intro]: "\\fin N x; a\N; x\M\ \ Cut .M (x).N \\<^sub>c N{x:=.M}" @@ -3098,7 +3098,7 @@ by (induct rule: c_redu.induct) (auto simp: abs_fresh rename_fresh subst_fresh)+ inductive - a_redu :: "trm \ trm \ bool" ("_ \\<^sub>a _" [100,100] 100) + a_redu :: "trm \ trm \ bool" (\_ \\<^sub>a _\ [100,100] 100) where al_redu[intro]: "M\\<^sub>l M' \ M \\<^sub>a M'" | ac_redu[intro]: "M\\<^sub>c M' \ M \\<^sub>a M'" @@ -3477,7 +3477,7 @@ text \Transitive Closure\ abbreviation - a_star_redu :: "trm \ trm \ bool" ("_ \\<^sub>a* _" [80,80] 80) + a_star_redu :: "trm \ trm \ bool" (\_ \\<^sub>a* _\ [80,80] 80) where "M \\<^sub>a* M' \ (a_redu)\<^sup>*\<^sup>* M M'" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Class2.thy --- a/src/HOL/Nominal/Examples/Class2.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Class2.thy Fri Sep 20 19:51:08 2024 +0200 @@ -2949,12 +2949,12 @@ done abbreviation - CANDn::"ty \ ntrm set" ("\'(_')\" [60] 60) + CANDn::"ty \ ntrm set" (\\'(_')\\ [60] 60) where "\(B)\ \ lfp (NEGn B \ NEGc B)" abbreviation - CANDc::"ty \ ctrm set" ("\<_>\" [60] 60) + CANDc::"ty \ ctrm set" (\\<_>\\ [60] 60) where "\\ \ NEGc B (\(B)\)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Class3.thy --- a/src/HOL/Nominal/Examples/Class3.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Class3.thy Fri Sep 20 19:51:08 2024 +0200 @@ -3143,12 +3143,12 @@ done abbreviation - nmaps :: "(name\coname\trm) list \ name \ (coname\trm) option \ bool" ("_ nmaps _ to _" [55,55,55] 55) + nmaps :: "(name\coname\trm) list \ name \ (coname\trm) option \ bool" (\_ nmaps _ to _\ [55,55,55] 55) where "\_n nmaps x to P \ (findn \_n x) = P" abbreviation - cmaps :: "(coname\name\trm) list \ coname \ (name\trm) option \ bool" ("_ cmaps _ to _" [55,55,55] 55) + cmaps :: "(coname\name\trm) list \ coname \ (name\trm) option \ bool" (\_ cmaps _ to _\ [55,55,55] 55) where "\_c cmaps a to P \ (findc \_c a) = P" @@ -3662,7 +3662,7 @@ done nominal_primrec (freshness_context: "(\_n::(name\coname\trm) list,\_c::(coname\name\trm) list)") - psubst :: "(name\coname\trm) list\(coname\name\trm) list\trm\trm" ("_,_<_>" [100,100,100] 100) + psubst :: "(name\coname\trm) list\(coname\name\trm) list\trm\trm" (\_,_<_>\ [100,100,100] 100) where "\_n,\_c = lookup x a \_n \_c" | "\a\(N,\_n,\_c);x\(M,\_n,\_c)\ \ \_n,\_c.M (x).N> = @@ -4863,12 +4863,12 @@ done definition - ncloses :: "(name\coname\trm) list\(name\ty) list \ bool" ("_ ncloses _" [55,55] 55) + ncloses :: "(name\coname\trm) list\(name\ty) list \ bool" (\_ ncloses _\ [55,55] 55) where "\_n ncloses \ \ \x B. ((x,B) \ set \ \ (\c P. \_n nmaps x to Some (c,P) \ :P \ (\\)))" definition - ccloses :: "(coname\name\trm) list\(coname\ty) list \ bool" ("_ ccloses _" [55,55] 55) + ccloses :: "(coname\name\trm) list\(coname\ty) list \ bool" (\_ ccloses _\ [55,55] 55) where "\_c ccloses \ \ \a B. ((a,B) \ set \ \ (\x P. \_c cmaps a to Some (x,P) \ (x):P \ (\(B)\)))" @@ -4938,7 +4938,7 @@ text \typing relation\ inductive - typing :: "ctxtn \ trm \ ctxtc \ bool" ("_ \ _ \ _" [100,100,100] 100) + typing :: "ctxtn \ trm \ ctxtc \ bool" (\_ \ _ \ _\ [100,100,100] 100) where TAx: "\validn \;validc \; (x,B)\set \; (a,B)\set \\ \ \ \ Ax x a \ \" | TNotR: "\x\\; ((x,B)#\) \ M \ \; set \' = {(a,NOT B)}\set \; validc \'\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Compile.thy --- a/src/HOL/Nominal/Examples/Compile.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Compile.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,11 +13,11 @@ nominal_datatype ty = Data "data" - | Arrow "ty" "ty" ("_\_" [100,100] 100) + | Arrow "ty" "ty" (\_\_\ [100,100] 100) nominal_datatype trm = Var "name" - | Lam "\name\trm" ("Lam [_]._" [100,100] 100) + | Lam "\name\trm" (\Lam [_]._\ [100,100] 100) | App "trm" "trm" | Const "nat" | Pr "trm" "trm" @@ -26,24 +26,24 @@ | InL "trm" | InR "trm" | Case "trm" "\name\trm" "\name\trm" - ("Case _ of inl _ \ _ | inr _ \ _" [100,100,100,100,100] 100) + (\Case _ of inl _ \ _ | inr _ \ _\ [100,100,100,100,100] 100) nominal_datatype dataI = OneI | NatI nominal_datatype tyI = DataI "dataI" - | ArrowI "tyI" "tyI" ("_\_" [100,100] 100) + | ArrowI "tyI" "tyI" (\_\_\ [100,100] 100) nominal_datatype trmI = IVar "name" - | ILam "\name\trmI" ("ILam [_]._" [100,100] 100) + | ILam "\name\trmI" (\ILam [_]._\ [100,100] 100) | IApp "trmI" "trmI" | IUnit | INat "nat" | ISucc "trmI" - | IAss "trmI" "trmI" ("_\_" [100,100] 100) + | IAss "trmI" "trmI" (\_\_\ [100,100] 100) | IRef "trmI" - | ISeq "trmI" "trmI" ("_;;_" [100,100] 100) + | ISeq "trmI" "trmI" (\_;;_\ [100,100] 100) | Iif "trmI" "trmI" "trmI" text \valid contexts\ @@ -57,7 +57,7 @@ text \typing judgements for trms\ inductive - typing :: "(name\ty) list\trm\ty\bool" (" _ \ _ : _ " [80,80,80] 80) + typing :: "(name\ty) list\trm\ty\bool" (\ _ \ _ : _ \ [80,80,80] 80) where t0[intro]: "\valid \; (x,\)\set \\\ \ \ Var x : \" | t1[intro]: "\\ \ e1 : \1\\2; \ \ e2 : \1\\ \ \ App e1 e2 : \2" @@ -75,7 +75,7 @@ text \typing judgements for Itrms\ inductive - Ityping :: "(name\tyI) list\trmI\tyI\bool" (" _ I\ _ : _ " [80,80,80] 80) + Ityping :: "(name\tyI) list\trmI\tyI\bool" (\ _ I\ _ : _ \ [80,80,80] 80) where t0[intro]: "\valid \; (x,\)\set \\\ \ I\ IVar x : \" | t1[intro]: "\\ I\ e1 : \1\\2; \ I\ e2 : \1\\ \ I\ IApp e1 e2 : \2" @@ -91,7 +91,7 @@ text \capture-avoiding substitution\ class subst = - fixes subst :: "'a \ name \ 'a \ 'a" ("_[_::=_]" [100,100,100] 100) + fixes subst :: "'a \ name \ 'a \ 'a" (\_[_::=_]\ [100,100,100] 100) instantiation trm :: subst begin @@ -169,7 +169,7 @@ text \big-step evaluation for trms\ inductive - big :: "trm\trm\bool" ("_ \ _" [80,80] 80) + big :: "trm\trm\bool" (\_ \ _\ [80,80] 80) where b0[intro]: "Lam [x].e \ Lam [x].e" | b1[intro]: "\e1\Lam [x].e; e2\e2'; e[x::=e2']\e'\ \ App e1 e2 \ e'" @@ -183,7 +183,7 @@ | b9[intro]: "\e\InR e'; e2[x::=e']\e''\ \ Case e of inl x1 \ e1 | inr x2 \ e2 \ e''" inductive - Ibig :: "((nat\nat)\trmI)\((nat\nat)\trmI)\bool" ("_ I\ _" [80,80] 80) + Ibig :: "((nat\nat)\trmI)\((nat\nat)\trmI)\bool" (\_ I\ _\ [80,80] 80) where m0[intro]: "(m,ILam [x].e) I\ (m,ILam [x].e)" | m1[intro]: "\(m,e1)I\(m',ILam [x].e); (m',e2)I\(m'',e3); (m'',e[x::=e3])I\(m''',e4)\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Contexts.thy --- a/src/HOL/Nominal/Examples/Contexts.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Contexts.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,7 +25,7 @@ nominal_datatype lam = Var "name" | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + | Lam "\name\lam" (\Lam [_]._\ [100,100] 100) text \ Contexts - the lambda case in contexts does not bind @@ -33,15 +33,15 @@ \ nominal_datatype ctx = - Hole ("\" 1000) + Hole (\\\ 1000) | CAppL "ctx" "lam" | CAppR "lam" "ctx" - | CLam "name" "ctx" ("CLam [_]._" [100,100] 100) + | CLam "name" "ctx" (\CLam [_]._\ [100,100] 100) text \Capture-Avoiding Substitution\ nominal_primrec - subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\ [100,100,100] 100) where "(Var x)[y::=s] = (if x=y then s else (Var x))" | "(App t\<^sub>1 t\<^sub>2)[y::=s] = App (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])" @@ -58,7 +58,7 @@ \ nominal_primrec - filling :: "ctx \ lam \ lam" ("_\_\" [100,100] 100) + filling :: "ctx \ lam \ lam" (\_\_\\ [100,100] 100) where "\\t\ = t" | "(CAppL E t')\t\ = App (E\t\) t'" @@ -79,7 +79,7 @@ text \The composition of two contexts.\ nominal_primrec - ctx_compose :: "ctx \ ctx \ ctx" ("_ \ _" [100,100] 100) + ctx_compose :: "ctx \ ctx \ ctx" (\_ \ _\ [100,100] 100) where "\ \ E' = E'" | "(CAppL E t') \ E' = CAppL (E \ E') t'" @@ -94,14 +94,14 @@ text \Beta-reduction via contexts in the Felleisen-Hieb style.\ inductive - ctx_red :: "lam\lam\bool" ("_ \x _" [80,80] 80) + ctx_red :: "lam\lam\bool" (\_ \x _\ [80,80] 80) where xbeta[intro]: "E\App (Lam [x].t) t'\ \x E\t[x::=t']\" text \Beta-reduction via congruence rules in the Plotkin style.\ inductive - cong_red :: "lam\lam\bool" ("_ \o _" [80,80] 80) + cong_red :: "lam\lam\bool" (\_ \o _\ [80,80] 80) where obeta[intro]: "App (Lam [x].t) t' \o t[x::=t']" | oapp1[intro]: "t \o t' \ App t t2 \o App t' t2" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Crary.thy --- a/src/HOL/Nominal/Examples/Crary.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Crary.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,13 +16,13 @@ nominal_datatype ty = TBase | TUnit - | Arrow "ty" "ty" ("_\_" [100,100] 100) + | Arrow "ty" "ty" (\_\_\ [100,100] 100) nominal_datatype trm = Unit - | Var "name" ("Var _" [100] 100) - | Lam "\name\trm" ("Lam [_]._" [100,100] 100) - | App "trm" "trm" ("App _ _" [110,110] 100) + | Var "name" (\Var _\ [100] 100) + | Lam "\name\trm" (\Lam [_]._\ [100,100] 100) + | App "trm" "trm" (\App _ _\ [110,110] 100) | Const "nat" type_synonym Ctxt = "(name\ty) list" @@ -94,7 +94,7 @@ (auto simp: fresh_list_cons fresh_prod fresh_atm) nominal_primrec - psubst :: "Subst \ trm \ trm" ("_<_>" [100,100] 130) + psubst :: "Subst \ trm \ trm" (\_<_>\ [100,100] 130) where "\<(Var x)> = (lookup \ x)" | "\<(App t\<^sub>1 t\<^sub>2)> = App \1> \2>" @@ -104,7 +104,7 @@ by(finite_guess | simp add: abs_fresh | fresh_guess)+ abbreviation - subst :: "trm \ name \ trm \ trm" ("_[_::=_]" [100,100,100] 100) + subst :: "trm \ name \ trm \ trm" (\_[_::=_]\ [100,100,100] 100) where "t[x::=t'] \ ([(x,t')])" @@ -268,7 +268,7 @@ valid_cons_elim_auto[elim]:"valid ((x,T)#\)" abbreviation - "sub_context" :: "Ctxt \ Ctxt \ bool" (" _ \ _ " [55,55] 55) + "sub_context" :: "Ctxt \ Ctxt \ bool" (\ _ \ _ \ [55,55] 55) where "\\<^sub>1 \ \\<^sub>2 \ \a T. (a,T)\set \\<^sub>1 \ (a,T)\set \\<^sub>2" @@ -298,7 +298,7 @@ (auto dest!: fresh_context) inductive - typing :: "Ctxt\trm\ty\bool" (" _ \ _ : _ " [60,60,60] 60) + typing :: "Ctxt\trm\ty\bool" (\ _ \ _ : _ \ [60,60,60] 60) where T_Var[intro]: "\valid \; (x,T)\set \\ \ \ \ Var x : T" | T_App[intro]: "\\ \ e\<^sub>1 : T\<^sub>1\T\<^sub>2; \ \ e\<^sub>2 : T\<^sub>1\ \ \ \ App e\<^sub>1 e\<^sub>2 : T\<^sub>2" @@ -334,7 +334,7 @@ section \Definitional Equivalence\ inductive - def_equiv :: "Ctxt\trm\trm\ty\bool" ("_ \ _ \ _ : _" [60,60] 60) + def_equiv :: "Ctxt\trm\trm\ty\bool" (\_ \ _ \ _ : _\ [60,60] 60) where Q_Refl[intro]: "\ \ t : T \ \ \ t \ t : T" | Q_Symm[intro]: "\ \ t \ s : T \ \ \ s \ t : T" @@ -360,7 +360,7 @@ section \Weak Head Reduction\ inductive - whr_def :: "trm\trm\bool" ("_ \ _" [80,80] 80) + whr_def :: "trm\trm\bool" (\_ \ _\ [80,80] 80) where QAR_Beta[intro]: "App (Lam [x]. t\<^sub>1) t\<^sub>2 \ t\<^sub>1[x::=t\<^sub>2]" | QAR_App[intro]: "t\<^sub>1 \ t\<^sub>1' \ App t\<^sub>1 t\<^sub>2 \ App t\<^sub>1' t\<^sub>2" @@ -387,12 +387,12 @@ section \Weak Head Normalisation\ abbreviation - nf :: "trm \ bool" ("_ \|" [100] 100) + nf :: "trm \ bool" (\_ \|\ [100] 100) where "t\| \ \(\ u. t \ u)" inductive - whn_def :: "trm\trm\bool" ("_ \ _" [80,80] 80) + whn_def :: "trm\trm\bool" (\_ \ _\ [80,80] 80) where QAN_Reduce[intro]: "\s \ t; t \ u\ \ s \ u" | QAN_Normal[intro]: "t\| \ t \ t" @@ -429,9 +429,9 @@ section \Algorithmic Term Equivalence and Algorithmic Path Equivalence\ inductive - alg_equiv :: "Ctxt\trm\trm\ty\bool" ("_ \ _ \ _ : _" [60,60,60,60] 60) + alg_equiv :: "Ctxt\trm\trm\ty\bool" (\_ \ _ \ _ : _\ [60,60,60,60] 60) and - alg_path_equiv :: "Ctxt\trm\trm\ty\bool" ("_ \ _ \ _ : _" [60,60,60,60] 60) + alg_path_equiv :: "Ctxt\trm\trm\ty\bool" (\_ \ _ \ _ : _\ [60,60,60,60] 60) where QAT_Base[intro]: "\s \ p; t \ q; \ \ p \ q : TBase\ \ \ \ s \ t : TBase" | QAT_Arrow[intro]: "\x\(\,s,t); (x,T\<^sub>1)#\ \ App s (Var x) \ App t (Var x) : T\<^sub>2\ @@ -592,7 +592,7 @@ section \Logical Equivalence\ -function log_equiv :: "(Ctxt \ trm \ trm \ ty \ bool)" ("_ \ _ is _ : _" [60,60,60,60] 60) +function log_equiv :: "(Ctxt \ trm \ trm \ ty \ bool)" (\_ \ _ is _ : _\ [60,60,60,60] 60) where "\ \ s is t : TUnit = True" | "\ \ s is t : TBase = \ \ s \ t : TBase" @@ -744,7 +744,7 @@ qed abbreviation - log_equiv_for_psubsts :: "Ctxt \ Subst \ Subst \ Ctxt \ bool" ("_ \ _ is _ over _" [60,60] 60) + log_equiv_for_psubsts :: "Ctxt \ Subst \ Subst \ Ctxt \ bool" (\_ \ _ is _ over _\ [60,60] 60) where "\' \ \ is \' over \ \ \x T. (x,T) \ set \ \ \' \ \ is \' : T" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Fsub.thy --- a/src/HOL/Nominal/Examples/Fsub.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Fsub.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,7 +17,7 @@ section \Types for Names, Nominal Datatype Declaration for Types and Terms\ no_syntax - "_Map" :: "maplets => 'a \ 'b" ("(1[_])") + "_Map" :: "maplets => 'a \ 'b" (\(1[_])\) text \The main point of this solution is to use names everywhere (be they bound, binding or free). In System \FSUB{} there are two kinds of names corresponding to @@ -36,15 +36,15 @@ nominal_datatype ty = Tvar "tyvrs" | Top - | Arrow "ty" "ty" (infixr "\" 200) + | Arrow "ty" "ty" (infixr \\\ 200) | Forall "\tyvrs\ty" "ty" nominal_datatype trm = Var "vrs" | Abs "\vrs\trm" "ty" | TAbs "\tyvrs\trm" "ty" - | App "trm" "trm" (infixl "\" 200) - | TApp "trm" "ty" (infixl "\\<^sub>\" 200) + | App "trm" "trm" (infixl \\\ 200) + | TApp "trm" "ty" (infixl \\\<^sub>\\ 200) text \To be polite to the eye, some more familiar notation is introduced. Because of the change in the order of arguments, one needs to use @@ -52,17 +52,17 @@ as given above for \<^term>\Arrow\.\ abbreviation - Forall_syn :: "tyvrs \ ty \ ty \ ty" ("(3\_<:_./ _)" [0, 0, 10] 10) + Forall_syn :: "tyvrs \ ty \ ty \ ty" (\(3\_<:_./ _)\ [0, 0, 10] 10) where "\X<:T\<^sub>1. T\<^sub>2 \ ty.Forall X T\<^sub>2 T\<^sub>1" abbreviation - Abs_syn :: "vrs \ ty \ trm \ trm" ("(3\_:_./ _)" [0, 0, 10] 10) + Abs_syn :: "vrs \ ty \ trm \ trm" (\(3\_:_./ _)\ [0, 0, 10] 10) where "\x:T. t \ trm.Abs x t T" abbreviation - TAbs_syn :: "tyvrs \ ty \ trm \ trm" ("(3\_<:_./ _)" [0, 0, 10] 10) + TAbs_syn :: "tyvrs \ ty \ trm \ trm" (\(3\_<:_./ _)\ [0, 0, 10] 10) where "\X<:T. t \ trm.TAbs X t T" @@ -223,7 +223,7 @@ in \<^term>\\\. The set of free variables of \<^term>\S\ is the \support\ of \<^term>\S\.\ -definition "closed_in" :: "ty \ env \ bool" ("_ closed'_in _" [100,100] 100) where +definition "closed_in" :: "ty \ env \ bool" (\_ closed'_in _\ [100,100] 100) where "S closed_in \ \ (supp S)\(ty_dom \)" lemma closed_in_eqvt[eqvt]: @@ -279,7 +279,7 @@ text \Now validity of a context is a straightforward inductive definition.\ inductive - valid_rel :: "env \ bool" ("\ _ ok" [100] 100) + valid_rel :: "env \ bool" (\\ _ ok\ [100] 100) where valid_nil[simp]: "\ [] ok" | valid_consT[simp]: "\\ \ ok; X\(ty_dom \); T closed_in \\ \ \ (TVarB X T#\) ok" @@ -386,7 +386,7 @@ by (finite_guess | fresh_guess | simp)+ nominal_primrec - subst_ty :: "ty \ tyvrs \ ty \ ty" ("_[_ \ _]\<^sub>\" [300, 0, 0] 300) + subst_ty :: "ty \ tyvrs \ ty \ ty" (\_[_ \ _]\<^sub>\\ [300, 0, 0] 300) where "(Tvar X)[Y \ T]\<^sub>\ = (if X=Y then T else Tvar X)" | "(Top)[Y \ T]\<^sub>\ = Top" @@ -439,7 +439,7 @@ (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux) nominal_primrec - subst_tyb :: "binding \ tyvrs \ ty \ binding" ("_[_ \ _]\<^sub>b" [100,100,100] 100) + subst_tyb :: "binding \ tyvrs \ ty \ binding" (\_[_ \ _]\<^sub>b\ [100,100,100] 100) where "(TVarB X U)[Y \ T]\<^sub>b = TVarB X (U[Y \ T]\<^sub>\)" | "(VarB X U)[Y \ T]\<^sub>b = VarB X (U[Y \ T]\<^sub>\)" @@ -459,7 +459,7 @@ by (induct B rule: binding.induct) (simp_all add: fresh_atm type_subst_identity) -primrec subst_tyc :: "env \ tyvrs \ ty \ env" ("_[_ \ _]\<^sub>e" [100,100,100] 100) where +primrec subst_tyc :: "env \ tyvrs \ ty \ env" (\_[_ \ _]\<^sub>e\ [100,100,100] 100) where "([])[Y \ T]\<^sub>e= []" | "(B#\)[Y \ T]\<^sub>e = (B[Y \ T]\<^sub>b)#(\[Y \ T]\<^sub>e)" @@ -485,7 +485,7 @@ by (induct \) simp_all nominal_primrec - subst_trm :: "trm \ vrs \ trm \ trm" ("_[_ \ _]" [300, 0, 0] 300) + subst_trm :: "trm \ vrs \ trm \ trm" (\_[_ \ _]\ [300, 0, 0] 300) where "(Var x)[y \ t'] = (if x=y then t' else (Var x))" | "(t1 \ t2)[y \ t'] = t1[y \ t'] \ t2[y \ t']" @@ -525,7 +525,7 @@ (simp_all add: fresh_atm calc_atm abs_fresh fresh_aux ty_vrs_fresh perm_fresh_fresh) nominal_primrec (freshness_context: "T2::ty") - subst_trm_ty :: "trm \ tyvrs \ ty \ trm" ("_[_ \\<^sub>\ _]" [300, 0, 0] 300) + subst_trm_ty :: "trm \ tyvrs \ ty \ trm" (\_[_ \\<^sub>\ _]\ [300, 0, 0] 300) where "(Var x)[Y \\<^sub>\ T2] = Var x" | "(t1 \ t2)[Y \\<^sub>\ T2] = t1[Y \\<^sub>\ T2] \ t2[Y \\<^sub>\ T2]" @@ -576,7 +576,7 @@ $\alpha$-equivalence classes.)\ inductive - subtype_of :: "env \ ty \ ty \ bool" ("_\_<:_" [100,100,100] 100) + subtype_of :: "env \ ty \ ty \ bool" (\_\_<:_\ [100,100,100] 100) where SA_Top[intro]: "\\ \ ok; S closed_in \\ \ \ \ S <: Top" | SA_refl_TVar[intro]: "\\ \ ok; X \ ty_dom \\\ \ \ Tvar X <: Tvar X" @@ -698,7 +698,7 @@ another. This generalization seems to make the proof for weakening to be smoother than if we had strictly adhered to the version in the POPLmark-paper.\ -definition extends :: "env \ env \ bool" ("_ extends _" [100,100] 100) where +definition extends :: "env \ env \ bool" (\_ extends _\ [100,100] 100) where "\ extends \ \ \X Q. (TVarB X Q)\set \ \ (TVarB X Q)\set \" lemma extends_ty_dom: @@ -1017,7 +1017,7 @@ section \Typing\ inductive - typing :: "env \ trm \ ty \ bool" ("_ \ _ : _" [60,60,60] 60) + typing :: "env \ trm \ ty \ bool" (\_ \ _ : _\ [60,60,60] 60) where T_Var[intro]: "\ VarB x T \ set \; \ \ ok \ \ \ \ Var x : T" | T_App[intro]: "\ \ \ t\<^sub>1 : T\<^sub>1 \ T\<^sub>2; \ \ t\<^sub>2 : T\<^sub>1 \ \ \ \ t\<^sub>1 \ t\<^sub>2 : T\<^sub>2" @@ -1145,7 +1145,7 @@ "val (t1 \\<^sub>\ t2)" inductive - eval :: "trm \ trm \ bool" ("_ \ _" [60,60] 60) + eval :: "trm \ trm \ bool" (\_ \ _\ [60,60] 60) where E_Abs : "\ x \ v\<^sub>2; val v\<^sub>2 \ \ (\x:T\<^sub>11. t\<^sub>12) \ v\<^sub>2 \ t\<^sub>12[x \ v\<^sub>2]" | E_App1 [intro]: "t \ t' \ t \ u \ t' \ u" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Height.thy --- a/src/HOL/Nominal/Examples/Height.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Height.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ nominal_datatype lam = Var "name" | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + | Lam "\name\lam" (\Lam [_]._\ [100,100] 100) text \Definition of the height-function on lambda-terms.\ @@ -31,7 +31,7 @@ text \Definition of capture-avoiding substitution.\ nominal_primrec - subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\ [100,100,100] 100) where "(Var x)[y::=t'] = (if x=y then t' else (Var x))" | "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Lam_Funs.thy --- a/src/HOL/Nominal/Examples/Lam_Funs.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Lam_Funs.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ nominal_datatype lam = Var "name" | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + | Lam "\name\lam" (\Lam [_]._\ [100,100] 100) text \The depth of a lambda-term.\ @@ -68,7 +68,7 @@ by (induct \) (auto simp add: eqvts) nominal_primrec - psubst :: "(name\lam) list \ lam \ lam" ("_<_>" [95,95] 105) + psubst :: "(name\lam) list \ lam \ lam" (\_<_>\ [95,95] 105) where "\<(Var x)> = (lookup \ x)" | "\<(App e\<^sub>1 e\<^sub>2)> = App (\1>) (\2>)" @@ -83,7 +83,7 @@ (simp_all add: eqvts fresh_bij) abbreviation - subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\ [100,100,100] 100) where "t[x::=t'] \ ([(x,t')])" @@ -115,15 +115,15 @@ name, even if we introduce the notation [_]._ for CLam. \ nominal_datatype clam = - Hole ("\" 1000) + Hole (\\\ 1000) | CAppL "clam" "lam" | CAppR "lam" "clam" - | CLam "name" "clam" ("CLam [_]._" [100,100] 100) + | CLam "name" "clam" (\CLam [_]._\ [100,100] 100) text \Filling a lambda-term into a context.\ nominal_primrec - filling :: "clam \ lam \ lam" ("_\_\" [100,100] 100) + filling :: "clam \ lam \ lam" (\_\_\\ [100,100] 100) where "\\t\ = t" | "(CAppL E t')\t\ = App (E\t\) t'" @@ -134,7 +134,7 @@ text \Composition od two contexts\ nominal_primrec - clam_compose :: "clam \ clam \ clam" ("_ \ _" [100,100] 100) + clam_compose :: "clam \ clam \ clam" (\_ \ _\ [100,100] 100) where "\ \ E' = E'" | "(CAppL E t') \ E' = CAppL (E \ E') t'" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Lambda_mu.thy --- a/src/HOL/Nominal/Examples/Lambda_mu.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Lambda_mu.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,10 +8,10 @@ nominal_datatype trm = Var "var" - | Lam "\var\trm" ("Lam [_]._" [100,100] 100) + | Lam "\var\trm" (\Lam [_]._\ [100,100] 100) | App "trm" "trm" | Pss "mvar" "trm" (* passivate *) - | Act "\mvar\trm" ("Act [_]._" [100,100] 100) (* activate *) + | Act "\mvar\trm" (\Act [_]._\ [100,100] 100) (* activate *) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/LocalWeakening.thy --- a/src/HOL/Nominal/Examples/LocalWeakening.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/LocalWeakening.thy Fri Sep 20 19:51:08 2024 +0200 @@ -70,7 +70,7 @@ nominal_datatype ty = TVar "nat" - | TArr "ty" "ty" (infix "\" 200) + | TArr "ty" "ty" (infix \\\ 200) lemma ty_fresh[simp]: fixes x::"name" @@ -99,7 +99,7 @@ text \"weak" typing relation\ inductive - typing :: "cxt\llam\ty\bool" (" _ \ _ : _ " [60,60,60] 60) + typing :: "cxt\llam\ty\bool" (\ _ \ _ : _ \ [60,60,60] 60) where t_lPar[intro]: "\valid \; (x,T)\set \\\ \ \ lPar x : T" | t_lApp[intro]: "\\ \ t1 : T1\T2; \ \ t2 : T1\\ \ \ lApp t1 t2 : T2" @@ -126,7 +126,7 @@ text \sub-contexts\ abbreviation - "sub_context" :: "cxt \ cxt \ bool" ("_ \ _" [60,60] 60) + "sub_context" :: "cxt \ cxt \ bool" (\_ \ _\ [60,60] 60) where "\1 \ \2 \ \x T. (x,T)\set \1 \ (x,T)\set \2" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Pattern.thy --- a/src/HOL/Nominal/Examples/Pattern.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Pattern.thy Fri Sep 20 19:51:08 2024 +0200 @@ -5,14 +5,14 @@ begin no_syntax - "_Map" :: "maplets => 'a \ 'b" ("(1[_])") + "_Map" :: "maplets => 'a \ 'b" (\(1[_])\) atom_decl name nominal_datatype ty = Atom nat - | Arrow ty ty (infixr "\" 200) - | TupleT ty ty (infixr "\" 210) + | Arrow ty ty (infixr \\\ 200) + | TupleT ty ty (infixr \\\ 210) lemma fresh_type [simp]: "(a::name) \ (T::ty)" by (induct T rule: ty.induct) (simp_all add: fresh_nat) @@ -25,22 +25,22 @@ nominal_datatype trm = Var name - | Tuple trm trm ("(1'\_,/ _'\)") + | Tuple trm trm (\(1'\_,/ _'\)\) | Abs ty "\name\trm" - | App trm trm (infixl "\" 200) + | App trm trm (infixl \\\ 200) | Let ty trm btrm and btrm = Base trm | Bind ty "\name\btrm" abbreviation - Abs_syn :: "name \ ty \ trm \ trm" ("(3\_:_./ _)" [0, 0, 10] 10) + Abs_syn :: "name \ ty \ trm \ trm" (\(3\_:_./ _)\ [0, 0, 10] 10) where "\x:T. t \ Abs T x t" datatype pat = PVar name ty - | PTuple pat pat ("(1'\\_,/ _'\\)") + | PTuple pat pat (\(1'\\_,/ _'\\)\) (* FIXME: The following should be done automatically by the nominal package *) overloading pat_perm \ "perm :: name prm \ pat \ pat" (unchecked) @@ -83,7 +83,7 @@ (* the following function cannot be defined using nominal_primrec, *) (* since variable parameters are currently not allowed. *) -primrec abs_pat :: "pat \ btrm \ btrm" ("(3\[_]./ _)" [0, 10] 10) +primrec abs_pat :: "pat \ btrm \ btrm" (\(3\[_]./ _)\ [0, 10] 10) where "(\[PVar x T]. t) = Bind T x t" | "(\[\\p, q\\]. t) = (\[p]. \[q]. t)" @@ -143,7 +143,7 @@ type_synonym ctx = "(name \ ty) list" inductive - ptyping :: "pat \ ty \ ctx \ bool" ("\ _ : _ \ _" [60, 60, 60] 60) + ptyping :: "pat \ ty \ ctx \ bool" (\\ _ : _ \ _\ [60, 60, 60] 60) where PVar: "\ PVar x T : T \ [(x, T)]" | PTuple: "\ p : T \ \\<^sub>1 \ \ q : U \ \\<^sub>2 \ \ \\p, q\\ : T \ U \ \\<^sub>2 @ \\<^sub>1" @@ -168,16 +168,16 @@ by (induct \) (auto simp add: fresh_ctxt_set_eq [symmetric]) abbreviation - "sub_ctx" :: "ctx \ ctx \ bool" ("_ \ _") + "sub_ctx" :: "ctx \ ctx \ bool" (\_ \ _\) where "\\<^sub>1 \ \\<^sub>2 \ \x. x \ set \\<^sub>1 \ x \ set \\<^sub>2" abbreviation - Let_syn :: "pat \ trm \ trm \ trm" ("(LET (_ =/ _)/ IN (_))" 10) + Let_syn :: "pat \ trm \ trm \ trm" (\(LET (_ =/ _)/ IN (_))\ 10) where "LET p = t IN u \ Let (pat_type p) t (\[p]. Base u)" -inductive typing :: "ctx \ trm \ ty \ bool" ("_ \ _ : _" [60, 60, 60] 60) +inductive typing :: "ctx \ trm \ ty \ bool" (\_ \ _ : _\ [60, 60, 60] 60) where Var [intro]: "valid \ \ (x, T) \ set \ \ \ \ Var x : T" | Tuple [intro]: "\ \ t : T \ \ \ u : U \ \ \ \t, u\ : T \ U" @@ -268,7 +268,7 @@ qed auto inductive - match :: "pat \ trm \ (name \ trm) list \ bool" ("\ _ \ _ \ _" [50, 50, 50] 50) + match :: "pat \ trm \ (name \ trm) list \ bool" (\\ _ \ _ \ _\ [50, 50, 50] 50) where PVar: "\ PVar x T \ t \ [(x, t)]" | PProd: "\ p \ t \ \ \ \ q \ u \ \' \ \ \\p, q\\ \ \t, u\ \ \ @ \'" @@ -287,8 +287,8 @@ by (induct \) (auto simp add: eqvts) nominal_primrec - psubst :: "(name \ trm) list \ trm \ trm" ("_\_\" [95,0] 210) - and psubstb :: "(name \ trm) list \ btrm \ btrm" ("_\_\\<^sub>b" [95,0] 210) + psubst :: "(name \ trm) list \ trm \ trm" (\_\_\\ [95,0] 210) + and psubstb :: "(name \ trm) list \ btrm \ btrm" (\_\_\\<^sub>b\ [95,0] 210) where "\\Var x\ = (lookup \ x)" | "\\t \ u\ = \\t\ \ \\u\" @@ -320,12 +320,12 @@ (simp_all add: eqvts fresh_bij) abbreviation - subst :: "trm \ name \ trm \ trm" ("_[_\_]" [100,0,0] 100) + subst :: "trm \ name \ trm \ trm" (\_[_\_]\ [100,0,0] 100) where "t[x\t'] \ [(x,t')]\t\" abbreviation - substb :: "btrm \ name \ trm \ btrm" ("_[_\_]\<^sub>b" [100,0,0] 100) + substb :: "btrm \ name \ trm \ btrm" (\_[_\_]\<^sub>b\ [100,0,0] 100) where "t[x\t']\<^sub>b \ [(x,t')]\t\\<^sub>b" @@ -525,7 +525,7 @@ lemmas match_type = match_type_aux [where \\<^sub>1="[]", simplified] -inductive eval :: "trm \ trm \ bool" ("_ \ _" [60,60] 60) +inductive eval :: "trm \ trm \ bool" (\_ \ _\ [60,60] 60) where TupleL: "t \ t' \ \t, u\ \ \t', u\" | TupleR: "u \ u' \ \t, u\ \ \t, u'\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/SN.thy --- a/src/HOL/Nominal/Examples/SN.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/SN.thy Fri Sep 20 19:51:08 2024 +0200 @@ -71,7 +71,7 @@ (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh') inductive - Beta :: "lam\lam\bool" (" _ \\<^sub>\ _" [80,80] 80) + Beta :: "lam\lam\bool" (\ _ \\<^sub>\ _\ [80,80] 80) where b1[intro!]: "s1 \\<^sub>\ s2 \ App s1 t \\<^sub>\ App s2 t" | b2[intro!]: "s1\\<^sub>\s2 \ App t s1 \\<^sub>\ App t s2" @@ -113,7 +113,7 @@ nominal_datatype ty = TVar "nat" - | TArr "ty" "ty" (infix "\" 200) + | TArr "ty" "ty" (infix \\\ 200) lemma fresh_ty: fixes a ::"name" @@ -144,7 +144,7 @@ (auto simp add: fresh_prod fresh_list_cons fresh_atm) inductive - typing :: "(name\ty) list\lam\ty\bool" ("_ \ _ : _" [60,60,60] 60) + typing :: "(name\ty) list\lam\ty\bool" (\_ \ _ : _\ [60,60,60] 60) where t1[intro]: "\valid \; (a,\)\set \\ \ \ \ Var a : \" | t2[intro]: "\\ \ t1 : \\\; \ \ t2 : \\ \ \ \ App t1 t2 : \" @@ -230,7 +230,7 @@ (* a slight hack to get the first element of applications *) (* this is needed to get (SN t) from SN (App t s) *) inductive - FST :: "lam\lam\bool" (" _ \ _" [80,80] 80) + FST :: "lam\lam\bool" (\ _ \ _\ [80,80] 80) where fst[intro!]: "(App t s) \ t" @@ -455,12 +455,12 @@ qed abbreviation - mapsto :: "(name\lam) list \ name \ lam \ bool" ("_ maps _ to _" [55,55,55] 55) + mapsto :: "(name\lam) list \ name \ lam \ bool" (\_ maps _ to _\ [55,55,55] 55) where "\ maps x to e \ (lookup \ x) = e" abbreviation - closes :: "(name\lam) list \ (name\ty) list \ bool" ("_ closes _" [55,55] 55) + closes :: "(name\lam) list \ (name\ty) list \ bool" (\_ closes _\ [55,55] 55) where "\ closes \ \ \x T. ((x,T) \ set \ \ (\t. \ maps x to t \ t \ RED T))" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/SOS.thy --- a/src/HOL/Nominal/Examples/SOS.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/SOS.thy Fri Sep 20 19:51:08 2024 +0200 @@ -19,11 +19,11 @@ text \types and terms\ nominal_datatype ty = TVar "nat" - | Arrow "ty" "ty" ("_\_" [100,100] 100) + | Arrow "ty" "ty" (\_\_\ [100,100] 100) nominal_datatype trm = Var "name" - | Lam "\name\trm" ("Lam [_]._" [100,100] 100) + | Lam "\name\trm" (\Lam [_]._\ [100,100] 100) | App "trm" "trm" lemma fresh_ty: @@ -62,7 +62,7 @@ (* parallel substitution *) nominal_primrec - psubst :: "(name\trm) list \ trm \ trm" ("_<_>" [95,95] 105) + psubst :: "(name\trm) list \ trm \ trm" (\_<_>\ [95,95] 105) where "\<(Var x)> = (lookup \ x)" | "\<(App e\<^sub>1 e\<^sub>2)> = App (\1>) (\2>)" @@ -92,7 +92,7 @@ text \Single substitution\ abbreviation - subst :: "trm \ name \ trm \ trm" ("_[_::=_]" [100,100,100] 100) + subst :: "trm \ name \ trm \ trm" (\_[_::=_]\ [100,100,100] 100) where "t[x::=t'] \ ([(x,t')])" @@ -157,7 +157,7 @@ text \Typing Relation\ inductive - typing :: "(name\ty) list\trm\ty\bool" ("_ \ _ : _" [60,60,60] 60) + typing :: "(name\ty) list\trm\ty\bool" (\_ \ _ : _\ [60,60,60] 60) where t_Var[intro]: "\valid \; (x,T)\set \\ \ \ \ Var x : T" | t_App[intro]: "\\ \ e\<^sub>1 : T\<^sub>1\T\<^sub>2; \ \ e\<^sub>2 : T\<^sub>1\ \ \ \ App e\<^sub>1 e\<^sub>2 : T\<^sub>2" @@ -188,7 +188,7 @@ (auto simp add: abs_fresh fresh_ty alpha trm.inject) abbreviation - "sub_context" :: "(name\ty) list \ (name\ty) list \ bool" ("_ \ _" [55,55] 55) + "sub_context" :: "(name\ty) list \ (name\ty) list \ bool" (\_ \ _\ [55,55] 55) where "\\<^sub>1 \ \\<^sub>2 \ \x T. (x,T)\set \\<^sub>1 \ (x,T)\set \\<^sub>2" @@ -257,7 +257,7 @@ text \Big-Step Evaluation\ inductive - big :: "trm\trm\bool" ("_ \ _" [80,80] 80) + big :: "trm\trm\bool" (\_ \ _\ [80,80] 80) where b_Lam[intro]: "Lam [x].e \ Lam [x].e" | b_App[intro]: "\x\(e\<^sub>1,e\<^sub>2,e'); e\<^sub>1\Lam [x].e; e\<^sub>2\e\<^sub>2'; e[x::=e\<^sub>2']\e'\ \ App e\<^sub>1 e\<^sub>2 \ e'" @@ -427,12 +427,12 @@ text \'\ maps x to e' asserts that \ substitutes x with e\ abbreviation - mapsto :: "(name\trm) list \ name \ trm \ bool" ("_ maps _ to _" [55,55,55] 55) + mapsto :: "(name\trm) list \ name \ trm \ bool" (\_ maps _ to _\ [55,55,55] 55) where "\ maps x to e \ (lookup \ x) = e" abbreviation - v_closes :: "(name\trm) list \ (name\ty) list \ bool" ("_ Vcloses _" [55,55] 55) + v_closes :: "(name\trm) list \ (name\ty) list \ bool" (\_ Vcloses _\ [55,55] 55) where "\ Vcloses \ \ \x T. (x,T) \ set \ \ (\v. \ maps x to v \ v \ V T)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Standardization.thy --- a/src/HOL/Nominal/Examples/Standardization.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Standardization.thy Fri Sep 20 19:51:08 2024 +0200 @@ -20,8 +20,8 @@ nominal_datatype lam = Var "name" -| App "lam" "lam" (infixl "\" 200) -| Lam "\name\lam" ("Lam [_]._" [0, 10] 10) +| App "lam" "lam" (infixl \\\ 200) +| Lam "\name\lam" (\Lam [_]._\ [0, 10] 10) instantiation lam :: size begin @@ -38,7 +38,7 @@ end nominal_primrec - subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [300, 0, 0] 300) + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\ [300, 0, 0] 300) where subst_Var: "(Var x)[y::=s] = (if x=y then s else (Var x))" | subst_App: "(t\<^sub>1 \ t\<^sub>2)[y::=s] = t\<^sub>1[y::=s] \ t\<^sub>2[y::=s]" @@ -82,7 +82,7 @@ lemma subst_neq [simp]: "x \ y \ (Var x)[y::=u] = Var x" by (simp add: subst_Var) -inductive beta :: "lam \ lam \ bool" (infixl "\\<^sub>\" 50) +inductive beta :: "lam \ lam \ bool" (infixl \\\<^sub>\\ 50) where beta: "x \ t \ (Lam [x].s) \ t \\<^sub>\ s[x::=t]" | appL [simp, intro!]: "s \\<^sub>\ t \ s \ u \\<^sub>\ t \ u" @@ -106,14 +106,14 @@ qed abbreviation - beta_reds :: "lam \ lam \ bool" (infixl "\\<^sub>\\<^sup>*" 50) where + beta_reds :: "lam \ lam \ bool" (infixl \\\<^sub>\\<^sup>*\ 50) where "s \\<^sub>\\<^sup>* t \ beta\<^sup>*\<^sup>* s t" subsection \Application of a term to a list of terms\ abbreviation - list_application :: "lam \ lam list \ lam" (infixl "\\" 150) where + list_application :: "lam \ lam list \ lam" (infixl \\\\ 150) where "t \\ ts \ foldl (\) t ts" lemma apps_eq_tail_conv [iff]: "(r \\ ts = s \\ ts) = (r = s)" @@ -318,7 +318,7 @@ subsection \Lifting beta-reduction to lists\ abbreviation - list_beta :: "lam list \ lam list \ bool" (infixl "[\\<^sub>\]\<^sub>1" 50) where + list_beta :: "lam list \ lam list \ bool" (infixl \[\\<^sub>\]\<^sub>1\ 50) where "rs [\\<^sub>\]\<^sub>1 ss \ step1 beta rs ss" lemma head_Var_reduction: @@ -400,8 +400,8 @@ by induct (simp_all add: listrelp.intros perm_app [symmetric]) inductive - sred :: "lam \ lam \ bool" (infixl "\\<^sub>s" 50) - and sredlist :: "lam list \ lam list \ bool" (infixl "[\\<^sub>s]" 50) + sred :: "lam \ lam \ bool" (infixl \\\<^sub>s\ 50) + and sredlist :: "lam list \ lam list \ bool" (infixl \[\\<^sub>s]\ 50) where "s [\\<^sub>s] t \ listrelp (\\<^sub>s) s t" | Var: "rs [\\<^sub>s] rs' \ Var x \\ rs \\<^sub>s Var x \\ rs'" @@ -697,8 +697,8 @@ subsection \Leftmost reduction and weakly normalizing terms\ inductive - lred :: "lam \ lam \ bool" (infixl "\\<^sub>l" 50) - and lredlist :: "lam list \ lam list \ bool" (infixl "[\\<^sub>l]" 50) + lred :: "lam \ lam \ bool" (infixl \\\<^sub>l\ 50) + and lredlist :: "lam list \ lam list \ bool" (infixl \[\\<^sub>l]\ 50) where "s [\\<^sub>l] t \ listrelp (\\<^sub>l) s t" | Var: "rs [\\<^sub>l] rs' \ Var x \\ rs \\<^sub>l Var x \\ rs'" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Type_Preservation.thy --- a/src/HOL/Nominal/Examples/Type_Preservation.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Type_Preservation.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,12 +15,12 @@ nominal_datatype lam = Var "name" | App "lam" "lam" -| Lam "\name\lam" ("Lam [_]._") +| Lam "\name\lam" (\Lam [_]._\) text \Capture-Avoiding Substitution\ nominal_primrec - subst :: "lam \ name \ lam \ lam" ("_[_::=_]") + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\) where "(Var x)[y::=s] = (if x=y then s else (Var x))" | "(App t\<^sub>1 t\<^sub>2)[y::=s] = App (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])" @@ -47,7 +47,7 @@ nominal_datatype ty = TVar "string" - | TArr "ty" "ty" ("_ \ _") + | TArr "ty" "ty" (\_ \ _\) lemma ty_fresh: fixes x::"name" @@ -61,7 +61,7 @@ type_synonym ctx = "(name\ty) list" abbreviation - "sub_ctx" :: "ctx \ ctx \ bool" ("_ \ _") + "sub_ctx" :: "ctx \ ctx \ bool" (\_ \ _\) where "\\<^sub>1 \ \\<^sub>2 \ \x. x \ set \\<^sub>1 \ x \ set \\<^sub>2" @@ -102,7 +102,7 @@ text \Typing Relation\ inductive - typing :: "ctx \ lam \ ty \ bool" ("_ \ _ : _") + typing :: "ctx \ lam \ ty \ bool" (\_ \ _ : _\) where t_Var[intro]: "\valid \; (x,T)\set \\ \ \ \ Var x : T" | t_App[intro]: "\\ \ t\<^sub>1 : T\<^sub>1\T\<^sub>2; \ \ t\<^sub>2 : T\<^sub>1\ \ \ \ App t\<^sub>1 t\<^sub>2 : T\<^sub>2" @@ -169,7 +169,7 @@ text \Beta Reduction\ inductive - "beta" :: "lam\lam\bool" (" _ \\<^sub>\ _") + "beta" :: "lam\lam\bool" (\ _ \\<^sub>\ _\) where b1[intro]: "t1 \\<^sub>\ t2 \ App t1 s \\<^sub>\ App t2 s" | b2[intro]: "s1 \\<^sub>\ s2 \ App t s1 \\<^sub>\ App t s2" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/VC_Condition.thy --- a/src/HOL/Nominal/Examples/VC_Condition.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/VC_Condition.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ nominal_datatype lam = Var "name" | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + | Lam "\name\lam" (\Lam [_]._\ [100,100] 100) text \ The inductive relation 'unbind' unbinds the top-most @@ -27,7 +27,7 @@ of the algorithm W.\ inductive - unbind :: "lam \ name list \ lam \ bool" ("_ \ _,_" [60,60,60] 60) + unbind :: "lam \ name list \ lam \ bool" (\_ \ _,_\ [60,60,60] 60) where u_var: "(Var a) \ [],(Var a)" | u_app: "(App t1 t2) \ [],(App t1 t2)" @@ -143,7 +143,7 @@ strips off the top-most binders from lambdas.\ inductive - strip :: "lam \ lam \ bool" ("_ \ _" [60,60] 60) + strip :: "lam \ lam \ bool" (\_ \ _\ [60,60] 60) where s_var: "(Var a) \ (Var a)" | s_app: "(App t1 t2) \ (App t1 t2)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/W.thy --- a/src/HOL/Nominal/Examples/W.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/W.thy Fri Sep 20 19:51:08 2024 +0200 @@ -7,7 +7,7 @@ atom_decl tvar var abbreviation - "difference_list" :: "'a list \ 'a list \ 'a list" ("_ - _" [60,60] 60) + "difference_list" :: "'a list \ 'a list \ 'a list" (\_ - _\ [60,60] 60) where "xs - ys \ [x \ xs. x\set ys]" @@ -34,20 +34,20 @@ nominal_datatype ty = TVar "tvar" - | Fun "ty" "ty" ("_\_" [100,100] 100) + | Fun "ty" "ty" (\_\_\ [100,100] 100) nominal_datatype tyS = Ty "ty" - | ALL "\tvar\tyS" ("\[_]._" [100,100] 100) + | ALL "\tvar\tyS" (\\[_]._\ [100,100] 100) nominal_datatype trm = Var "var" | App "trm" "trm" - | Lam "\var\trm" ("Lam [_]._" [100,100] 100) + | Lam "\var\trm" (\Lam [_]._\ [100,100] 100) | Let "\var\trm" "trm" abbreviation - LetBe :: "var \ trm \ trm \ trm" ("Let _ be _ in _" [100,100,100] 100) + LetBe :: "var \ trm \ trm \ trm" (\Let _ be _ in _\ [100,100,100] 100) where "Let x be t1 in t2 \ trm.Let x t2 t1" @@ -175,10 +175,10 @@ type_synonym Subst = "(tvar\ty) list" consts - psubst :: "Subst \ 'a \ 'a" ("_<_>" [100,60] 120) + psubst :: "Subst \ 'a \ 'a" (\_<_>\ [100,60] 120) abbreviation - subst :: "'a \ tvar \ ty \ 'a" ("_[_::=_]" [100,100,100] 100) + subst :: "'a \ tvar \ ty \ 'a" (\_[_::=_]\ [100,100,100] 100) where "smth[X::=T] \ ([(X,T)])" @@ -299,7 +299,7 @@ text \instance of a type scheme\ inductive - inst :: "ty \ tyS \ bool"("_ \ _" [50,51] 50) + inst :: "ty \ tyS \ bool"(\_ \ _\ [50,51] 50) where I_Ty[intro]: "T \ (Ty T)" | I_All[intro]: "\X\T'; T \ S\ \ T[X::=T'] \ \[X].S" @@ -350,7 +350,7 @@ text\typing judgements\ inductive - typing :: "Ctxt \ trm \ ty \ bool" (" _ \ _ : _ " [60,60,60] 60) + typing :: "Ctxt \ trm \ ty \ bool" (\ _ \ _ : _ \ [60,60,60] 60) where T_VAR[intro]: "\valid \; (x,S)\set \; T \ S\\ \ \ Var x : T" | T_APP[intro]: "\\ \ t\<^sub>1 : T\<^sub>1\T\<^sub>2; \ \ t\<^sub>2 : T\<^sub>1\\ \ \ App t\<^sub>1 t\<^sub>2 : T\<^sub>2" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Weakening.thy --- a/src/HOL/Nominal/Examples/Weakening.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Weakening.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,11 +14,11 @@ nominal_datatype lam = Var "name" | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + | Lam "\name\lam" (\Lam [_]._\ [100,100] 100) nominal_datatype ty = TVar "string" - | TArr "ty" "ty" ("_ \ _" [100,100] 100) + | TArr "ty" "ty" (\_ \ _\ [100,100] 100) lemma ty_fresh: fixes x::"name" @@ -42,7 +42,7 @@ text\Typing judgements\ inductive - typing :: "(name\ty) list\lam\ty\bool" ("_ \ _ : _" [60,60,60] 60) + typing :: "(name\ty) list\lam\ty\bool" (\_ \ _ : _\ [60,60,60] 60) where t_Var[intro]: "\valid \; (x,T)\set \\ \ \ \ Var x : T" | t_App[intro]: "\\ \ t1 : T1\T2; \ \ t2 : T1\ \ \ \ App t1 t2 : T2" @@ -60,7 +60,7 @@ text \Abbreviation for the notion of subcontexts.\ abbreviation - "sub_context" :: "(name\ty) list \ (name\ty) list \ bool" ("_ \ _" [60,60] 60) + "sub_context" :: "(name\ty) list \ (name\ty) list \ bool" (\_ \ _\ [60,60] 60) where "\1 \ \2 \ \x T. (x,T)\set \1 \ (x,T)\set \2" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nonstandard_Analysis/HDeriv.thy --- a/src/HOL/Nonstandard_Analysis/HDeriv.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nonstandard_Analysis/HDeriv.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,12 +13,12 @@ text \Nonstandard Definitions.\ definition nsderiv :: "['a::real_normed_field \ 'a, 'a, 'a] \ bool" - ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) + (\(NSDERIV (_)/ (_)/ :> (_))\ [1000, 1000, 60] 60) where "NSDERIV f x :> D \ (\h \ Infinitesimal - {0}. (( *f* f)(star_of x + h) - star_of (f x)) / h \ star_of D)" definition NSdifferentiable :: "['a::real_normed_field \ 'a, 'a] \ bool" - (infixl "NSdifferentiable" 60) + (infixl \NSdifferentiable\ 60) where "f NSdifferentiable x \ (\D. NSDERIV f x :> D)" definition increment :: "(real \ real) \ real \ hypreal \ hypreal" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nonstandard_Analysis/HLim.thy --- a/src/HOL/Nonstandard_Analysis/HLim.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nonstandard_Analysis/HLim.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ text \Nonstandard Definitions.\ definition NSLIM :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ 'b \ bool" - ("((_)/ \(_)/\\<^sub>N\<^sub>S (_))" [60, 0, 60] 60) + (\((_)/ \(_)/\\<^sub>N\<^sub>S (_))\ [60, 0, 60] 60) where "f \a\\<^sub>N\<^sub>S L \ (\x. x \ star_of a \ x \ star_of a \ ( *f* f) x \ star_of L)" definition isNSCont :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ bool" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nonstandard_Analysis/HLog.thy --- a/src/HOL/Nonstandard_Analysis/HLog.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nonstandard_Analysis/HLog.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,7 +9,7 @@ imports HTranscendental begin -definition powhr :: "hypreal \ hypreal \ hypreal" (infixr "powhr" 80) +definition powhr :: "hypreal \ hypreal \ hypreal" (infixr \powhr\ 80) where [transfer_unfold]: "x powhr a = starfun2 (powr) x a" definition hlog :: "hypreal \ hypreal \ hypreal" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nonstandard_Analysis/HSEQ.thy --- a/src/HOL/Nonstandard_Analysis/HSEQ.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nonstandard_Analysis/HSEQ.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ begin definition NSLIMSEQ :: "(nat \ 'a::real_normed_vector) \ 'a \ bool" - ("((_)/ \\<^sub>N\<^sub>S (_))" [60, 60] 60) where + (\((_)/ \\<^sub>N\<^sub>S (_))\ [60, 60] 60) where \ \Nonstandard definition of convergence of sequence\ "X \\<^sub>N\<^sub>S L \ (\N \ HNatInfinite. ( *f* X) N \ star_of L)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nonstandard_Analysis/HSeries.thy --- a/src/HOL/Nonstandard_Analysis/HSeries.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nonstandard_Analysis/HSeries.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,7 +14,7 @@ definition sumhr :: "hypnat \ hypnat \ (nat \ real) \ hypreal" where "sumhr = (\(M,N,f). starfun2 (\m n. sum f {m.. real) \ real \ bool" (infixr "NSsums" 80) +definition NSsums :: "(nat \ real) \ real \ bool" (infixr \NSsums\ 80) where "f NSsums s = (\n. sum f {..\<^sub>N\<^sub>S s" definition NSsummable :: "(nat \ real) \ bool" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nonstandard_Analysis/HyperDef.thy --- a/src/HOL/Nonstandard_Analysis/HyperDef.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nonstandard_Analysis/HyperDef.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,11 +18,11 @@ abbreviation hypreal_of_hypnat :: "hypnat \ hypreal" where "hypreal_of_hypnat \ of_hypnat" -definition omega :: hypreal ("\") +definition omega :: hypreal (\\\) where "\ = star_n (\n. real (Suc n))" \ \an infinite number \= [<1, 2, 3, \>]\\ -definition epsilon :: hypreal ("\") +definition epsilon :: hypreal (\\\) where "\ = star_n (\n. inverse (real (Suc n)))" \ \an infinitesimal number \= [<1, 1/2, 1/3, \>]\\ @@ -301,7 +301,7 @@ subsection \Powers with Hypernatural Exponents\ text \Hypernatural powers of hyperreals.\ -definition pow :: "'a::power star \ nat star \ 'a star" (infixr "pow" 80) +definition pow :: "'a::power star \ nat star \ 'a star" (infixr \pow\ 80) where hyperpow_def [transfer_unfold]: "R pow N = ( *f2* (^)) R N" lemma Standard_hyperpow [simp]: "r \ Standard \ n \ Standard \ r pow n \ Standard" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nonstandard_Analysis/NSA.thy --- a/src/HOL/Nonstandard_Analysis/NSA.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nonstandard_Analysis/NSA.thy Fri Sep 20 19:51:08 2024 +0200 @@ -21,7 +21,7 @@ definition HInfinite :: "('a::real_normed_vector) star set" where "HInfinite = {x. \r \ Reals. r < hnorm x}" -definition approx :: "'a::real_normed_vector star \ 'a star \ bool" (infixl "\" 50) +definition approx :: "'a::real_normed_vector star \ 'a star \ bool" (infixl \\\ 50) where "x \ y \ x - y \ Infinitesimal" \ \the ``infinitely close'' relation\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nonstandard_Analysis/Star.thy --- a/src/HOL/Nonstandard_Analysis/Star.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nonstandard_Analysis/Star.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ begin definition \ \internal sets\ - starset_n :: "(nat \ 'a set) \ 'a star set" ("*sn* _" [80] 80) + starset_n :: "(nat \ 'a set) \ 'a star set" (\*sn* _\ [80] 80) where "*sn* As = Iset (star_n As)" definition InternalSets :: "'a star set set" @@ -23,7 +23,7 @@ (\x y. \X \ Rep_star x. \Y \ Rep_star y. y = F x \ eventually (\n. Y n = f(X n)) \)" definition \ \internal functions\ - starfun_n :: "(nat \ 'a \ 'b) \ 'a star \ 'b star" ("*fn* _" [80] 80) + starfun_n :: "(nat \ 'a \ 'b) \ 'a star \ 'b star" (\*fn* _\ [80] 80) where "*fn* F = Ifun (star_n F)" definition InternalFuns :: "('a star => 'b star) set" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Number_Theory/Eratosthenes.thy --- a/src/HOL/Number_Theory/Eratosthenes.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Number_Theory/Eratosthenes.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,7 +14,7 @@ context dvd begin -abbreviation dvd_strict :: "'a \ 'a \ bool" (infixl "dvd'_strict" 50) +abbreviation dvd_strict :: "'a \ 'a \ bool" (infixl \dvd'_strict\ 50) where "b dvd_strict a \ b dvd a \ \ a dvd b" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Predicate_Compile_Examples/Examples.thy --- a/src/HOL/Predicate_Compile_Examples/Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Predicate_Compile_Examples/Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -166,15 +166,15 @@ datatype type = Atom nat - | Fun type type (infixr "\" 200) + | Fun type type (infixr \\\ 200) datatype dB = Var nat - | App dB dB (infixl "\" 200) + | App dB dB (infixl \\\ 200) | Abs type dB primrec - nth_el :: "'a list \ nat \ 'a option" ("_\_\" [90, 0] 91) + nth_el :: "'a list \ nat \ 'a option" (\_\_\\ [90, 0] 91) where "[]\i\ = None" | "(x # xs)\i\ = (case i of 0 \ Some x | Suc j \ xs \j\)" @@ -184,7 +184,7 @@ "nth_el' (x # xs) 0 x" | "nth_el' xs i y \ nth_el' (x # xs) (Suc i) y" -inductive typing :: "type list \ dB \ type \ bool" ("_ \ _ : _" [50, 50, 50] 50) +inductive typing :: "type list \ dB \ type \ bool" (\_ \ _ : _\ [50, 50, 50] 50) where Var [intro!]: "nth_el' env x T \ env \ Var x : T" | Abs [intro!]: "T # env \ t : U \ env \ Abs T t : (T \ U)" @@ -198,14 +198,14 @@ | "lift (Abs T s) k = Abs T (lift s (k + 1))" primrec - subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300) + subst :: "[dB, dB, nat] => dB" (\_[_'/_]\ [300, 0, 0] 300) where subst_Var: "(Var i)[s/k] = (if k < i then Var (i - 1) else if i = k then s else Var i)" | subst_App: "(t \ u)[s/k] = t[s/k] \ u[s/k]" | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])" -inductive beta :: "[dB, dB] => bool" (infixl "\\<^sub>\" 50) +inductive beta :: "[dB, dB] => bool" (infixl \\\<^sub>\\ 50) where beta [simp, intro!]: "Abs T s \ t \\<^sub>\ s[t/0]" | appL [simp, intro!]: "s \\<^sub>\ t ==> s \ u \\<^sub>\ t \ u" @@ -284,13 +284,13 @@ datatype com' = SKIP - | Assign name aexp ("_ ::= _" [1000, 61] 61) - | Semi com' com' ("_; _" [60, 61] 60) - | If bexp com' com' ("IF _ THEN _ ELSE _" [0, 0, 61] 61) - | While bexp com' ("WHILE _ DO _" [0, 61] 61) + | Assign name aexp (\_ ::= _\ [1000, 61] 61) + | Semi com' com' (\_; _\ [60, 61] 60) + | If bexp com' com' (\IF _ THEN _ ELSE _\ [0, 0, 61] 61) + | While bexp com' (\WHILE _ DO _\ [0, 61] 61) inductive - big_step :: "com' * state' \ state' \ bool" (infix "\" 55) + big_step :: "com' * state' \ state' \ bool" (infix \\\ 55) where Skip: "(SKIP,s) \ s" | Assign: "(x ::= a,s) \ s(x := aval a s)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Predicate_Compile_Examples/Lambda_Example.thy --- a/src/HOL/Predicate_Compile_Examples/Lambda_Example.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Predicate_Compile_Examples/Lambda_Example.thy Fri Sep 20 19:51:08 2024 +0200 @@ -6,15 +6,15 @@ datatype type = Atom nat - | Fun type type (infixr "\" 200) + | Fun type type (infixr \\\ 200) datatype dB = Var nat - | App dB dB (infixl "\" 200) + | App dB dB (infixl \\\ 200) | Abs type dB primrec - nth_el :: "'a list \ nat \ 'a option" ("_\_\" [90, 0] 91) + nth_el :: "'a list \ nat \ 'a option" (\_\_\\ [90, 0] 91) where "[]\i\ = None" | "(x # xs)\i\ = (case i of 0 \ Some x | Suc j \ xs \j\)" @@ -24,7 +24,7 @@ "nth_el1 (x # xs) 0 x" | "nth_el1 xs i y \ nth_el1 (x # xs) (Suc i) y" -inductive typing :: "type list \ dB \ type \ bool" ("_ \ _ : _" [50, 50, 50] 50) +inductive typing :: "type list \ dB \ type \ bool" (\_ \ _ : _\ [50, 50, 50] 50) where Var [intro!]: "nth_el1 env x T \ env \ Var x : T" | Abs [intro!]: "T # env \ t : U \ env \ Abs T t : (T \ U)" @@ -42,14 +42,14 @@ "pred (Suc i) = i" primrec - subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300) + subst :: "[dB, dB, nat] => dB" (\_[_'/_]\ [300, 0, 0] 300) where subst_Var: "(Var i)[s/k] = (if k < i then Var (pred i) else if i = k then s else Var i)" | subst_App: "(t \ u)[s/k] = t[s/k] \ u[s/k]" | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])" -inductive beta :: "[dB, dB] => bool" (infixl "\\<^sub>\" 50) +inductive beta :: "[dB, dB] => bool" (infixl \\\<^sub>\\ 50) where beta [simp, intro!]: "Abs T s \ t \\<^sub>\ s[t/0]" | appL [simp, intro!]: "s \\<^sub>\ t ==> s \ u \\<^sub>\ t \ u" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy --- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -210,15 +210,15 @@ datatype type = Atom nat - | Fun type type (infixr "\" 200) + | Fun type type (infixr \\\ 200) datatype dB = Var nat - | App dB dB (infixl "\" 200) + | App dB dB (infixl \\\ 200) | Abs type dB primrec - nth_el :: "'a list \ nat \ 'a option" ("_\_\" [90, 0] 91) + nth_el :: "'a list \ nat \ 'a option" (\_\_\\ [90, 0] 91) where "[]\i\ = None" | "(x # xs)\i\ = (case i of 0 \ Some x | Suc j \ xs \j\)" @@ -228,7 +228,7 @@ "nth_el' (x # xs) 0 x" | "nth_el' xs i y \ nth_el' (x # xs) (Suc i) y" -inductive typing :: "type list \ dB \ type \ bool" ("_ \ _ : _" [50, 50, 50] 50) +inductive typing :: "type list \ dB \ type \ bool" (\_ \ _ : _\ [50, 50, 50] 50) where Var [intro!]: "nth_el' env x T \ env \ Var x : T" | Abs [intro!]: "T # env \ t : U \ env \ Abs T t : (T \ U)" @@ -242,14 +242,14 @@ | "lift (Abs T s) k = Abs T (lift s (k + 1))" primrec - subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300) + subst :: "[dB, dB, nat] => dB" (\_[_'/_]\ [300, 0, 0] 300) where subst_Var: "(Var i)[s/k] = (if k < i then Var (i - 1) else if i = k then s else Var i)" | subst_App: "(t \ u)[s/k] = t[s/k] \ u[s/k]" | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])" -inductive beta :: "[dB, dB] => bool" (infixl "\\<^sub>\" 50) +inductive beta :: "[dB, dB] => bool" (infixl \\\<^sub>\\ 50) where beta [simp, intro!]: "Abs T s \ t \\<^sub>\ s[t/0]" | appL [simp, intro!]: "s \\<^sub>\ t ==> s \ u \\<^sub>\ t \ u" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy --- a/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Predicate_Compile_Examples/Specialisation_Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -95,24 +95,24 @@ section \Specialisation in POPLmark theory\ notation - Some ("\_\") + Some (\\_\\) notation - None ("\") + None (\\\) notation - length ("\_\") + length (\\_\\) notation - Cons ("_ \/ _" [66, 65] 65) + Cons (\_ \/ _\ [66, 65] 65) primrec - nth_el :: "'a list \ nat \ 'a option" ("_\_\" [90, 0] 91) + nth_el :: "'a list \ nat \ 'a option" (\_\_\\ [90, 0] 91) where "[]\i\ = \" | "(x # xs)\i\ = (case i of 0 \ \x\ | Suc j \ xs \j\)" -primrec assoc :: "('a \ 'b) list \ 'a \ 'b option" ("_\_\\<^sub>?" [90, 0] 91) +primrec assoc :: "('a \ 'b) list \ 'a \ 'b option" (\_\_\\<^sub>?\ [90, 0] 91) where "[]\a\\<^sub>? = \" | "(x # xs)\a\\<^sub>? = (if fst x = a then \snd x\ else xs\a\\<^sub>?)" @@ -125,8 +125,8 @@ datatype type = TVar nat | Top - | Fun type type (infixr "\" 200) - | TyAll type type ("(3\<:_./ _)" [0, 10] 10) + | Fun type type (infixr \\\ 200) + | TyAll type type (\(3\<:_./ _)\ [0, 10] 10) datatype binding = VarB type | TVarB type type_synonym env = "binding list" @@ -148,19 +148,19 @@ datatype trm = Var nat - | Abs type trm ("(3\:_./ _)" [0, 10] 10) - | TAbs type trm ("(3\<:_./ _)" [0, 10] 10) - | App trm trm (infixl "\" 200) - | TApp trm type (infixl "\\<^sub>\" 200) + | Abs type trm (\(3\:_./ _)\ [0, 10] 10) + | TAbs type trm (\(3\<:_./ _)\ [0, 10] 10) + | App trm trm (infixl \\\ 200) + | TApp trm type (infixl \\\<^sub>\\ 200) -primrec liftT :: "nat \ nat \ type \ type" ("\\<^sub>\") +primrec liftT :: "nat \ nat \ type \ type" (\\\<^sub>\\) where "\\<^sub>\ n k (TVar i) = (if i < k then TVar i else TVar (i + n))" | "\\<^sub>\ n k Top = Top" | "\\<^sub>\ n k (T \ U) = \\<^sub>\ n k T \ \\<^sub>\ n k U" | "\\<^sub>\ n k (\<:T. U) = (\<:\\<^sub>\ n k T. \\<^sub>\ n (k + 1) U)" -primrec lift :: "nat \ nat \ trm \ trm" ("\") +primrec lift :: "nat \ nat \ trm \ trm" (\\\) where "\ n k (Var i) = (if i < k then Var i else Var (i + n))" | "\ n k (\:T. t) = (\:\\<^sub>\ n k T. \ n (k + 1) t)" @@ -168,7 +168,7 @@ | "\ n k (s \ t) = \ n k s \ \ n k t" | "\ n k (t \\<^sub>\ T) = \ n k t \\<^sub>\ \\<^sub>\ n k T" -primrec substTT :: "type \ nat \ type \ type" ("_[_ \\<^sub>\ _]\<^sub>\" [300, 0, 0] 300) +primrec substTT :: "type \ nat \ type \ type" (\_[_ \\<^sub>\ _]\<^sub>\\ [300, 0, 0] 300) where "(TVar i)[k \\<^sub>\ S]\<^sub>\ = (if k < i then TVar (i - 1) else if i = k then \\<^sub>\ k 0 S else TVar i)" @@ -176,12 +176,12 @@ | "(T \ U)[k \\<^sub>\ S]\<^sub>\ = T[k \\<^sub>\ S]\<^sub>\ \ U[k \\<^sub>\ S]\<^sub>\" | "(\<:T. U)[k \\<^sub>\ S]\<^sub>\ = (\<:T[k \\<^sub>\ S]\<^sub>\. U[k+1 \\<^sub>\ S]\<^sub>\)" -primrec decT :: "nat \ nat \ type \ type" ("\\<^sub>\") +primrec decT :: "nat \ nat \ type \ type" (\\\<^sub>\\) where "\\<^sub>\ 0 k T = T" | "\\<^sub>\ (Suc n) k T = \\<^sub>\ n k (T[k \\<^sub>\ Top]\<^sub>\)" -primrec subst :: "trm \ nat \ trm \ trm" ("_[_ \ _]" [300, 0, 0] 300) +primrec subst :: "trm \ nat \ trm \ trm" (\_[_ \ _]\ [300, 0, 0] 300) where "(Var i)[k \ s] = (if k < i then Var (i - 1) else if i = k then \ k 0 s else Var i)" | "(t \ u)[k \ s] = t[k \ s] \ u[k \ s]" @@ -189,7 +189,7 @@ | "(\:T. t)[k \ s] = (\:\\<^sub>\ 1 k T. t[k+1 \ s])" | "(\<:T. t)[k \ s] = (\<:\\<^sub>\ 1 k T. t[k+1 \ s])" -primrec substT :: "trm \ nat \ type \ trm" ("_[_ \\<^sub>\ _]" [300, 0, 0] 300) +primrec substT :: "trm \ nat \ type \ trm" (\_[_ \\<^sub>\ _]\ [300, 0, 0] 300) where "(Var i)[k \\<^sub>\ S] = (if k < i then Var (i - 1) else Var i)" | "(t \ u)[k \\<^sub>\ S] = t[k \\<^sub>\ S] \ u[k \\<^sub>\ S]" @@ -197,23 +197,23 @@ | "(\:T. t)[k \\<^sub>\ S] = (\:T[k \\<^sub>\ S]\<^sub>\. t[k+1 \\<^sub>\ S])" | "(\<:T. t)[k \\<^sub>\ S] = (\<:T[k \\<^sub>\ S]\<^sub>\. t[k+1 \\<^sub>\ S])" -primrec liftE :: "nat \ nat \ env \ env" ("\\<^sub>e") +primrec liftE :: "nat \ nat \ env \ env" (\\\<^sub>e\) where "\\<^sub>e n k [] = []" | "\\<^sub>e n k (B \ \) = mapB (\\<^sub>\ n (k + \\\)) B \ \\<^sub>e n k \" -primrec substE :: "env \ nat \ type \ env" ("_[_ \\<^sub>\ _]\<^sub>e" [300, 0, 0] 300) +primrec substE :: "env \ nat \ type \ env" (\_[_ \\<^sub>\ _]\<^sub>e\ [300, 0, 0] 300) where "[][k \\<^sub>\ T]\<^sub>e = []" | "(B \ \)[k \\<^sub>\ T]\<^sub>e = mapB (\U. U[k + \\\ \\<^sub>\ T]\<^sub>\) B \ \[k \\<^sub>\ T]\<^sub>e" -primrec decE :: "nat \ nat \ env \ env" ("\\<^sub>e") +primrec decE :: "nat \ nat \ env \ env" (\\\<^sub>e\) where "\\<^sub>e 0 k \ = \" | "\\<^sub>e (Suc n) k \ = \\<^sub>e n k (\[k \\<^sub>\ Top]\<^sub>e)" inductive - well_formed :: "env \ type \ bool" ("_ \\<^sub>w\<^sub>f _" [50, 50] 50) + well_formed :: "env \ type \ bool" (\_ \\<^sub>w\<^sub>f _\ [50, 50] 50) where wf_TVar: "\\i\ = \TVarB T\ \ \ \\<^sub>w\<^sub>f TVar i" | wf_Top: "\ \\<^sub>w\<^sub>f Top" @@ -221,8 +221,8 @@ | wf_all: "\ \\<^sub>w\<^sub>f T \ TVarB T \ \ \\<^sub>w\<^sub>f U \ \ \\<^sub>w\<^sub>f (\<:T. U)" inductive - well_formedE :: "env \ bool" ("_ \\<^sub>w\<^sub>f" [50] 50) - and well_formedB :: "env \ binding \ bool" ("_ \\<^sub>w\<^sub>f\<^sub>B _" [50, 50] 50) + well_formedE :: "env \ bool" (\_ \\<^sub>w\<^sub>f\ [50] 50) + and well_formedB :: "env \ binding \ bool" (\_ \\<^sub>w\<^sub>f\<^sub>B _\ [50, 50] 50) where "\ \\<^sub>w\<^sub>f\<^sub>B B \ \ \\<^sub>w\<^sub>f type_ofB B" | wf_Nil: "[] \\<^sub>w\<^sub>f" @@ -238,7 +238,7 @@ "B \ \ \\<^sub>w\<^sub>f" inductive - subtyping :: "env \ type \ type \ bool" ("_ \ _ <: _" [50, 50, 50] 50) + subtyping :: "env \ type \ type \ bool" (\_ \ _ <: _\ [50, 50, 50] 50) where SA_Top: "\ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f S \ \ \ S <: Top" | SA_refl_TVar: "\ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f TVar i \ \ \ TVar i <: TVar i" @@ -249,7 +249,7 @@ \ \ (\<:S\<^sub>1. S\<^sub>2) <: (\<:T\<^sub>1. T\<^sub>2)" inductive - typing :: "env \ trm \ type \ bool" ("_ \ _ : _" [50, 50, 50] 50) + typing :: "env \ trm \ type \ bool" (\_ \ _ : _\ [50, 50, 50] 50) where T_Var: "\ \\<^sub>w\<^sub>f \ \\i\ = \VarB U\ \ T = \\<^sub>\ (Suc i) 0 U \ \ \ Var i : T" | T_Abs: "VarB T\<^sub>1 \ \ \ t\<^sub>2 : T\<^sub>2 \ \ \ (\:T\<^sub>1. t\<^sub>2) : T\<^sub>1 \ \\<^sub>\ 1 0 T\<^sub>2" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/Convolution.thy --- a/src/HOL/Probability/Convolution.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/Convolution.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ lemma (in finite_measure) sigma_finite_measure: "sigma_finite_measure M" .. -definition convolution :: "('a :: ordered_euclidean_space) measure \ 'a measure \ 'a measure" (infix "\" 50) where +definition convolution :: "('a :: ordered_euclidean_space) measure \ 'a measure \ 'a measure" (infix \\\ 50) where "convolution M N = distr (M \\<^sub>M N) borel (\(x, y). x + y)" lemma diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/Fin_Map.thy --- a/src/HOL/Probability/Fin_Map.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/Fin_Map.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ stay close to the developments of (finite) products \<^const>\Pi\<^sub>E\ and their sigma-algebra \<^const>\Pi\<^sub>M\.\ -type_notation fmap ("(_ \\<^sub>F /_)" [22, 21] 21) +type_notation fmap (\(_ \\<^sub>F /_)\ [22, 21] 21) unbundle fmap.lifting @@ -25,7 +25,7 @@ lemma finite_domain[simp, intro]: "finite (domain P)" by transfer simp -lift_definition proj :: "('i \\<^sub>F 'a) \ 'i \ 'a" ("'((_)')\<^sub>F" [0] 1000) is +lift_definition proj :: "('i \\<^sub>F 'a) \ 'i \ 'a" (\'((_)')\<^sub>F\ [0] 1000) is "\f x. if x \ dom f then the (f x) else undefined" . declare [[coercion proj]] @@ -89,7 +89,7 @@ "Pi' I A = { P. domain P = I \ (\i. i \ I \ (P)\<^sub>F i \ A i) } " syntax - "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\'' _\_./ _)" 10) + "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" (\(3\'' _\_./ _)\ 10) syntax_consts "_Pi'" == Pi' translations @@ -636,7 +636,7 @@ "Pi\<^sub>F I M \ PiF I M" syntax - "_PiF" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" ("(3\\<^sub>F _\_./ _)" 10) + "_PiF" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" (\(3\\<^sub>F _\_./ _)\ 10) syntax_consts "_PiF" == PiF translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/Information.thy Fri Sep 20 19:51:08 2024 +0200 @@ -446,7 +446,7 @@ qed abbreviation (in information_space) - mutual_information_Pow ("\'(_ ; _')") where + mutual_information_Pow (\\'(_ ; _')\) where "\(X ; Y) \ mutual_information b (count_space (X`space M)) (count_space (Y`space M)) X Y" lemma (in information_space) @@ -723,7 +723,7 @@ "entropy b S X = - KL_divergence b S (distr M S X)" abbreviation (in information_space) - entropy_Pow ("\'(_')") where + entropy_Pow (\\'(_')\) where "\(X) \ entropy b (count_space (X`space M)) X" lemma (in prob_space) distributed_RN_deriv: @@ -888,7 +888,7 @@ mutual_information b MX MZ X Z" abbreviation (in information_space) - conditional_mutual_information_Pow ("\'( _ ; _ | _ ')") where + conditional_mutual_information_Pow (\\'( _ ; _ | _ ')\) where "\(X ; Y | Z) \ conditional_mutual_information b (count_space (X ` space M)) (count_space (Y ` space M)) (count_space (Z ` space M)) X Y Z" @@ -1420,7 +1420,7 @@ enn2real (RN_deriv T (distr M T Y) y)) \distr M (S \\<^sub>M T) (\x. (X x, Y x)))" abbreviation (in information_space) - conditional_entropy_Pow ("\'(_ | _')") where + conditional_entropy_Pow (\\'(_ | _')\) where "\(X | Y) \ conditional_entropy b (count_space (X`space M)) (count_space (Y`space M)) X Y" lemma (in information_space) conditional_entropy_generic_eq: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/Probability_Mass_Function.thy --- a/src/HOL/Probability/Probability_Mass_Function.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/Probability_Mass_Function.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ begin text \Conflicting notation from \<^theory>\HOL-Analysis.Infinite_Sum\\ -no_notation Infinite_Sum.abs_summable_on (infixr "abs'_summable'_on" 46) +no_notation Infinite_Sum.abs_summable_on (infixr \abs'_summable'_on\ 46) lemma AE_emeasure_singleton: assumes x: "emeasure M {x} \ 0" and ae: "AE x in M. P x" shows "P x" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/Product_PMF.thy --- a/src/HOL/Probability/Product_PMF.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/Product_PMF.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,7 +8,7 @@ begin text \Conflicting notation from \<^theory>\HOL-Analysis.Infinite_Sum\\ -no_notation Infinite_Sum.abs_summable_on (infixr "abs'_summable'_on" 46) +no_notation Infinite_Sum.abs_summable_on (infixr \abs'_summable'_on\ 46) subsection \Preliminaries\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/Projective_Family.thy --- a/src/HOL/Probability/Projective_Family.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/Projective_Family.thy Fri Sep 20 19:51:08 2024 +0200 @@ -116,7 +116,7 @@ qed qed (auto elim!: generator.cases) -definition mu_G ("\G") where +definition mu_G (\\G\) where "\G A = (THE x. \J\I. finite J \ (\X\sets (Pi\<^sub>M J M). A = emb I J X \ x = emeasure (P J) X))" definition lim :: "('i \ 'a) measure" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/Projective_Limit.thy --- a/src/HOL/Probability/Projective_Limit.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/Projective_Limit.thy Fri Sep 20 19:51:08 2024 +0200 @@ -106,7 +106,7 @@ assumes X: "J \ I" "finite J" "X \ sets (\\<^sub>M i\J. borel)" shows "lim (emb I J X) = P J X" proof (rule emeasure_lim) - write mu_G ("\G") + write mu_G (\\G\) interpret generator: algebra "space (PiM I (\i. borel))" generator by (rule algebra_generator) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/SPMF.thy --- a/src/HOL/Probability/SPMF.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/SPMF.thy Fri Sep 20 19:51:08 2024 +0200 @@ -992,7 +992,7 @@ where "ord_spmf ord \ rel_pmf (ord_option ord)" locale ord_spmf_syntax begin -notation ord_spmf (infix "\\" 60) +notation ord_spmf (infix \\\\ 60) end lemma ord_spmf_map_spmf1: "ord_spmf R (map_spmf f p) = ord_spmf (\x. R (f x)) p" @@ -1805,7 +1805,7 @@ subsection \Restrictions on spmfs\ -definition restrict_spmf :: "'a spmf \ 'a set \ 'a spmf" (infixl "\" 110) +definition restrict_spmf :: "'a spmf \ 'a set \ 'a spmf" (infixl \\\ 110) where "p \ A = map_pmf (\x. x \ (\y. if y \ A then Some y else None)) p" lemma set_restrict_spmf [simp]: "set_spmf (p \ A) = set_spmf p \ A" @@ -2621,7 +2621,7 @@ subsection \Try\ -definition try_spmf :: "'a spmf \ 'a spmf \ 'a spmf" ("TRY _ ELSE _" [0,60] 59) +definition try_spmf :: "'a spmf \ 'a spmf \ 'a spmf" (\TRY _ ELSE _\ [0,60] 59) where "try_spmf p q = bind_pmf p (\x. case x of None \ q | Some y \ return_spmf y)" lemma try_spmf_lossless [simp]: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/ex/Dining_Cryptographers.thy --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Fri Sep 20 19:51:08 2024 +0200 @@ -402,13 +402,13 @@ by standard auto notation (in dining_cryptographers_space) - mutual_information_Pow ("\'( _ ; _ ')") + mutual_information_Pow (\\'( _ ; _ ')\) notation (in dining_cryptographers_space) - entropy_Pow ("\'( _ ')") + entropy_Pow (\\'( _ ')\) notation (in dining_cryptographers_space) - conditional_entropy_Pow ("\'( _ | _ ')") + conditional_entropy_Pow (\\'( _ | _ ')\) theorem (in dining_cryptographers_space) "\( inversion ; payer ) = 0" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Fri Sep 20 19:51:08 2024 +0200 @@ -213,25 +213,25 @@ abbreviation "\ \ point_measure msgs P" -abbreviation probability ("\

'(_') _") where +abbreviation probability (\\

'(_') _\) where "\

(X) x \ measure \ (X -` x \ msgs)" -abbreviation joint_probability ("\

'(_ ; _') _") where +abbreviation joint_probability (\\

'(_ ; _') _\) where "\

(X ; Y) x \ \

(\x. (X x, Y x)) x" -no_notation disj (infixr "|" 30) +no_notation disj (infixr \|\ 30) -abbreviation conditional_probability ("\

'(_ | _') _") where +abbreviation conditional_probability (\\

'(_ | _') _\) where "\

(X | Y) x \ (\

(X ; Y) x) / \

(Y) (snd`x)" notation - entropy_Pow ("\'( _ ')") + entropy_Pow (\\'( _ ')\) notation - conditional_entropy_Pow ("\'( _ | _ ')") + conditional_entropy_Pow (\\'( _ | _ ')\) notation - mutual_information_Pow ("\'( _ ; _ ')") + mutual_information_Pow (\\'( _ ; _ ')\) lemma t_eq_imp_bij_func: assumes "t xs = t ys" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Prolog/Func.thy --- a/src/HOL/Prolog/Func.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Prolog/Func.thy Fri Sep 20 19:51:08 2024 +0200 @@ -19,15 +19,15 @@ true :: tm and false :: tm and - "and" :: "tm \ tm \ tm" (infixr "and" 999) and - eq :: "tm \ tm \ tm" (infixr "eq" 999) and + "and" :: "tm \ tm \ tm" (infixr \and\ 999) and + eq :: "tm \ tm \ tm" (infixr \eq\ 999) and - Z :: tm ("Z") and + Z :: tm (\Z\) and S :: "tm \ tm" and - plus :: "tm \ tm \ tm" (infixl "+" 65) and - minus :: "tm \ tm \ tm" (infixl "-" 65) and - times :: "tm \ tm \ tm" (infixl "*" 70) and + plus :: "tm \ tm \ tm" (infixl \+\ 65) and + minus :: "tm \ tm \ tm" (infixl \-\ 65) and + times :: "tm \ tm \ tm" (infixl \*\ 70) and eval :: "tm \ tm \ bool" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Prolog/HOHH.thy --- a/src/HOL/Prolog/HOHH.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Prolog/HOHH.thy Fri Sep 20 19:51:08 2024 +0200 @@ -21,13 +21,13 @@ consts (* D-formulas (programs): D ::= !x. D | D .. D | D :- G | A *) - Dand :: "[bool, bool] => bool" (infixr ".." 28) - Dif :: "[bool, bool] => bool" (infixl ":-" 29) + Dand :: "[bool, bool] => bool" (infixr \..\ 28) + Dif :: "[bool, bool] => bool" (infixl \:-\ 29) (* G-formulas (goals): G ::= A | G & G | G | G | ? x. G | True | !x. G | D => G *) (*Dand' :: "[bool, bool] => bool" (infixr "," 35)*) - Dimp :: "[bool, bool] => bool" (infixr "=>" 27) + Dimp :: "[bool, bool] => bool" (infixr \=>\ 27) translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Prolog/Test.thy --- a/src/HOL/Prolog/Test.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Prolog/Test.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,16 +13,16 @@ typedecl 'a list consts - Nil :: "'a list" ("[]") - Cons :: "'a => 'a list => 'a list" (infixr "#" 65) + Nil :: "'a list" (\[]\) + Cons :: "'a => 'a list => 'a list" (infixr \#\ 65) text \List enumeration\ nonterminal list_args syntax - "" :: "'a \ list_args" ("_") - "_list_args" :: "'a \ list_args \ list_args" ("_,/ _") - "_list" :: "list_args => 'a list" ("[(_)]") + "" :: "'a \ list_args" (\_\) + "_list_args" :: "'a \ list_args \ list_args" (\_,/ _\) + "_list" :: "list_args => 'a list" (\[(_)]\) syntax_consts "_list_args" "_list" == Cons translations @@ -42,9 +42,9 @@ sue :: person and ned :: person and - nat23 :: nat ("23") and - nat24 :: nat ("24") and - nat25 :: nat ("25") and + nat23 :: nat (\23\) and + nat24 :: nat (\24\) and + nat25 :: nat (\25\) and age :: "[person, nat] => bool" and diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Prolog/Type.thy --- a/src/HOL/Prolog/Type.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Prolog/Type.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ axiomatization bool :: ty and nat :: ty and - arrow :: "ty \ ty \ ty" (infixr "->" 20) and + arrow :: "ty \ ty \ ty" (infixr \->\ 20) and typeof :: "[tm, ty] \ bool" and anyterm :: tm where common_typeof: " diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Proofs/Lambda/Eta.thy --- a/src/HOL/Proofs/Lambda/Eta.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Proofs/Lambda/Eta.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,7 +18,7 @@ | "free (Abs s) i = free s (i + 1)" inductive - eta :: "[dB, dB] => bool" (infixl "\\<^sub>\" 50) + eta :: "[dB, dB] => bool" (infixl \\\<^sub>\\ 50) where eta [simp, intro]: "\ free s 0 ==> Abs (s \ Var 0) \\<^sub>\ s[dummy/0]" | appL [simp, intro]: "s \\<^sub>\ t ==> s \ u \\<^sub>\ t \ u" @@ -26,11 +26,11 @@ | abs [simp, intro]: "s \\<^sub>\ t ==> Abs s \\<^sub>\ Abs t" abbreviation - eta_reds :: "[dB, dB] => bool" (infixl "\\<^sub>\\<^sup>*" 50) where + eta_reds :: "[dB, dB] => bool" (infixl \\\<^sub>\\<^sup>*\ 50) where "s \\<^sub>\\<^sup>* t == eta\<^sup>*\<^sup>* s t" abbreviation - eta_red0 :: "[dB, dB] => bool" (infixl "\\<^sub>\\<^sup>=" 50) where + eta_red0 :: "[dB, dB] => bool" (infixl \\\<^sub>\\<^sup>=\ 50) where "s \\<^sub>\\<^sup>= t == eta\<^sup>=\<^sup>= s t" inductive_cases eta_cases [elim!]: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Proofs/Lambda/Lambda.thy --- a/src/HOL/Proofs/Lambda/Lambda.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Proofs/Lambda/Lambda.thy Fri Sep 20 19:51:08 2024 +0200 @@ -16,7 +16,7 @@ datatype dB = Var nat - | App dB dB (infixl "\" 200) + | App dB dB (infixl \\\ 200) | Abs dB primrec @@ -27,7 +27,7 @@ | "lift (Abs s) k = Abs (lift s (k + 1))" primrec - subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300) + subst :: "[dB, dB, nat] => dB" (\_[_'/_]\ [300, 0, 0] 300) where (* FIXME base names *) subst_Var: "(Var i)[s/k] = (if k < i then Var (i - 1) else if i = k then s else Var i)" @@ -56,7 +56,7 @@ subsection \Beta-reduction\ -inductive beta :: "[dB, dB] => bool" (infixl "\\<^sub>\" 50) +inductive beta :: "[dB, dB] => bool" (infixl \\\<^sub>\\ 50) where beta [simp, intro!]: "Abs s \ t \\<^sub>\ s[t/0]" | appL [simp, intro!]: "s \\<^sub>\ t ==> s \ u \\<^sub>\ t \ u" @@ -64,7 +64,7 @@ | abs [simp, intro!]: "s \\<^sub>\ t ==> Abs s \\<^sub>\ Abs t" abbreviation - beta_reds :: "[dB, dB] => bool" (infixl "\\<^sub>\\<^sup>*" 50) where + beta_reds :: "[dB, dB] => bool" (infixl \\\<^sub>\\<^sup>*\ 50) where "s \\<^sub>\\<^sup>* t == beta\<^sup>*\<^sup>* s t" inductive_cases beta_cases [elim!]: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Proofs/Lambda/LambdaType.thy --- a/src/HOL/Proofs/Lambda/LambdaType.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Proofs/Lambda/LambdaType.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ subsection \Environments\ definition - shift :: "(nat \ 'a) \ nat \ 'a \ nat \ 'a" ("_\_:_\" [90, 0, 0] 91) where + shift :: "(nat \ 'a) \ nat \ 'a \ nat \ 'a" (\_\_:_\\ [90, 0, 0] 91) where "e\i:a\ = (\j. if j < i then e j else if j = i then a else e (j - 1))" lemma shift_eq [simp]: "i = j \ (e\i:T\) j = T" @@ -31,9 +31,9 @@ datatype type = Atom nat - | Fun type type (infixr "\" 200) + | Fun type type (infixr \\\ 200) -inductive typing :: "(nat \ type) \ dB \ type \ bool" ("_ \ _ : _" [50, 50, 50] 50) +inductive typing :: "(nat \ type) \ dB \ type \ bool" (\_ \ _ : _\ [50, 50, 50] 50) where Var [intro!]: "env x = T \ env \ Var x : T" | Abs [intro!]: "env\0:T\ \ t : U \ env \ Abs t : (T \ U)" @@ -55,11 +55,11 @@ abbreviation typings_rel :: "(nat \ type) \ dB list \ type list \ bool" - ("_ \ _ : _" [50, 50, 50] 50) where + (\_ \ _ : _\ [50, 50, 50] 50) where "env \ ts : Ts == typings env ts Ts" abbreviation - funs :: "type list \ type \ type" (infixr "\" 200) where + funs :: "type list \ type \ type" (infixr \\\ 200) where "Ts \ T == foldr Fun Ts T" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Proofs/Lambda/ListApplication.thy --- a/src/HOL/Proofs/Lambda/ListApplication.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Proofs/Lambda/ListApplication.thy Fri Sep 20 19:51:08 2024 +0200 @@ -8,7 +8,7 @@ theory ListApplication imports Lambda begin abbreviation - list_application :: "dB => dB list => dB" (infixl "\\" 150) where + list_application :: "dB => dB list => dB" (infixl \\\\ 150) where "t \\ ts == foldl (\) t ts" lemma apps_eq_tail_conv [iff]: "(r \\ ts = s \\ ts) = (r = s)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Proofs/Lambda/ListBeta.thy --- a/src/HOL/Proofs/Lambda/ListBeta.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Proofs/Lambda/ListBeta.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ \ abbreviation - list_beta :: "dB list => dB list => bool" (infixl "=>" 50) where + list_beta :: "dB list => dB list => bool" (infixl \=>\ 50) where "rs => ss == step1 beta rs ss" lemma head_Var_reduction: diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Proofs/Lambda/ParRed.thy --- a/src/HOL/Proofs/Lambda/ParRed.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Proofs/Lambda/ParRed.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ subsection \Parallel reduction\ -inductive par_beta :: "[dB, dB] => bool" (infixl "=>" 50) +inductive par_beta :: "[dB, dB] => bool" (infixl \=>\ 50) where var [simp, intro!]: "Var n => Var n" | abs [simp, intro!]: "s => t ==> Abs s => Abs t" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Proofs/Lambda/Standardization.thy --- a/src/HOL/Proofs/Lambda/Standardization.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Proofs/Lambda/Standardization.thy Fri Sep 20 19:51:08 2024 +0200 @@ -20,8 +20,8 @@ declare listrel_mono [mono_set] inductive - sred :: "dB \ dB \ bool" (infixl "\\<^sub>s" 50) - and sredlist :: "dB list \ dB list \ bool" (infixl "[\\<^sub>s]" 50) + sred :: "dB \ dB \ bool" (infixl \\\<^sub>s\ 50) + and sredlist :: "dB list \ dB list \ bool" (infixl \[\\<^sub>s]\ 50) where "s [\\<^sub>s] t \ listrelp (\\<^sub>s) s t" | Var: "rs [\\<^sub>s] rs' \ Var x \\ rs \\<^sub>s Var x \\ rs'" @@ -243,8 +243,8 @@ subsection \Leftmost reduction and weakly normalizing terms\ inductive - lred :: "dB \ dB \ bool" (infixl "\\<^sub>l" 50) - and lredlist :: "dB list \ dB list \ bool" (infixl "[\\<^sub>l]" 50) + lred :: "dB \ dB \ bool" (infixl \\\<^sub>l\ 50) + and lredlist :: "dB list \ dB list \ bool" (infixl \[\\<^sub>l]\ 50) where "s [\\<^sub>l] t \ listrelp (\\<^sub>l) s t" | Var: "rs [\\<^sub>l] rs' \ Var x \\ rs \\<^sub>l Var x \\ rs'" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Proofs/Lambda/WeakNorm.thy --- a/src/HOL/Proofs/Lambda/WeakNorm.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Proofs/Lambda/WeakNorm.thy Fri Sep 20 19:51:08 2024 +0200 @@ -201,7 +201,7 @@ \ \A computationally relevant copy of @{term "e \ t : T"}\ -inductive rtyping :: "(nat \ type) \ dB \ type \ bool" ("_ \\<^sub>R _ : _" [50, 50, 50] 50) +inductive rtyping :: "(nat \ type) \ dB \ type \ bool" (\_ \\<^sub>R _ : _\ [50, 50, 50] 50) where Var: "e x = T \ e \\<^sub>R Var x : T" | Abs: "e\0:T\ \\<^sub>R t : U \ e \\<^sub>R Abs t : (T \ U)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Quickcheck_Examples/Quickcheck_Lattice_Examples.thy --- a/src/HOL/Quickcheck_Examples/Quickcheck_Lattice_Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Quickcheck_Examples/Quickcheck_Lattice_Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,12 +13,12 @@ the standard default type \<^typ>\int\ is insufficient.\ notation - less_eq (infix "\" 50) and - less (infix "\" 50) and - top ("\") and - bot ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) + less_eq (infix \\\ 50) and + less (infix \\\ 50) and + top (\\\) and + bot (\\\) and + inf (infixl \\\ 70) and + sup (infixl \\\ 65) declare [[quickcheck_narrowing_active = false, quickcheck_timeout = 3600]] @@ -133,11 +133,11 @@ no_notation - less_eq (infix "\" 50) and - less (infix "\" 50) and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - top ("\") and - bot ("\") + less_eq (infix \\\ 50) and + less (infix \\\ 50) and + inf (infixl \\\ 70) and + sup (infixl \\\ 65) and + top (\\\) and + bot (\\\) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Quotient_Examples/Lift_FSet.thy --- a/src/HOL/Quotient_Examples/Lift_FSet.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Quotient_Examples/Lift_FSet.thy Fri Sep 20 19:51:08 2024 +0200 @@ -38,7 +38,7 @@ subsection \Lifted constant definitions\ -lift_definition fnil :: "'a fset" ("{||}") is "[]" parametric list.ctr_transfer(1) . +lift_definition fnil :: "'a fset" (\{||}\) is "[]" parametric list.ctr_transfer(1) . lift_definition fcons :: "'a \ 'a fset \ 'a fset" is Cons parametric list.ctr_transfer(2) by simp @@ -97,19 +97,19 @@ nonterminal fset_args syntax - "" :: "'a \ fset_args" ("_") - "_fset_args" :: "'a \ fset_args \ fset_args" ("_,/ _") - "_fset" :: "fset_args => 'a fset" ("{|(_)|}") + "" :: "'a \ fset_args" (\_\) + "_fset_args" :: "'a \ fset_args \ fset_args" (\_,/ _\) + "_fset" :: "fset_args => 'a fset" (\{|(_)|}\) syntax_consts "_fset_args" "_fset" == fcons translations "{|x, xs|}" == "CONST fcons x {|xs|}" "{|x|}" == "CONST fcons x {||}" -lift_definition fmember :: "'a \ 'a fset \ bool" (infix "|\|" 50) is "\x xs. x \ set xs" +lift_definition fmember :: "'a \ 'a fset \ bool" (infix \|\|\ 50) is "\x xs. x \ set xs" parametric member_transfer by simp -abbreviation notin_fset :: "'a \ 'a fset \ bool" (infix "|\|" 50) where +abbreviation notin_fset :: "'a \ 'a fset \ bool" (infix \|\|\ 50) where "x |\| S \ \ (x |\| S)" lemma fmember_fmap[simp]: "a |\| fmap f X = (\b. b |\| X \ a = f b)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Quotient_Examples/Lift_Fun.thy --- a/src/HOL/Quotient_Examples/Lift_Fun.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Quotient_Examples/Lift_Fun.thy Fri Sep 20 19:51:08 2024 +0200 @@ -19,7 +19,7 @@ in the following four definitions. This example is based on HOL/Fun.thy.\ quotient_type -('a, 'b) fun' (infixr "\" 55) = "'a \ 'b" / "(=)" +('a, 'b) fun' (infixr \\\ 55) = "'a \ 'b" / "(=)" by (simp add: identity_equivp) quotient_definition "comp' :: ('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" is diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Quotient_Examples/Quotient_FSet.thy --- a/src/HOL/Quotient_Examples/Quotient_FSet.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Quotient_Examples/Quotient_FSet.thy Fri Sep 20 19:51:08 2024 +0200 @@ -22,7 +22,7 @@ \ definition - list_eq :: "'a list \ 'a list \ bool" (infix "\" 50) + list_eq :: "'a list \ 'a list \ bool" (infix \\\ 50) where [simp]: "xs \ ys \ set xs = set ys" @@ -199,7 +199,7 @@ is "Nil :: 'a list" done abbreviation - empty_fset ("{||}") + empty_fset (\{||}\) where "{||} \ bot :: 'a fset" @@ -208,7 +208,7 @@ is "sub_list :: ('a list \ 'a list \ bool)" by simp abbreviation - subset_fset :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) + subset_fset :: "'a fset \ 'a fset \ bool" (infix \|\|\ 50) where "xs |\| ys \ xs \ ys" @@ -218,7 +218,7 @@ "xs < ys \ xs \ ys \ xs \ (ys::'a fset)" abbreviation - psubset_fset :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) + psubset_fset :: "'a fset \ 'a fset \ bool" (infix \|\|\ 50) where "xs |\| ys \ xs < ys" @@ -227,7 +227,7 @@ is "append :: 'a list \ 'a list \ 'a list" by simp abbreviation - union_fset (infixl "|\|" 65) + union_fset (infixl \|\|\ 65) where "xs |\| ys \ sup xs (ys::'a fset)" @@ -236,7 +236,7 @@ is "inter_list :: 'a list \ 'a list \ 'a list" by simp abbreviation - inter_fset (infixl "|\|" 65) + inter_fset (infixl \|\|\ 65) where "xs |\| ys \ inf xs (ys::'a fset)" @@ -290,9 +290,9 @@ nonterminal fset_args syntax - "" :: "'a \ fset_args" ("_") - "_fset_args" :: "'a \ fset_args \ fset_args" ("_,/ _") - "_fset" :: "fset_args => 'a fset" ("{|(_)|}") + "" :: "'a \ fset_args" (\_\) + "_fset_args" :: "'a \ fset_args \ fset_args" (\_,/ _\) + "_fset" :: "fset_args => 'a fset" (\{|(_)|}\) syntax_consts "_fset_args" "_fset" == insert_fset translations @@ -305,12 +305,12 @@ "fset_member :: 'a fset \ 'a \ bool" is "List.member" by fastforce abbreviation - in_fset :: "'a \ 'a fset \ bool" (infix "|\|" 50) + in_fset :: "'a \ 'a fset \ bool" (infix \|\|\ 50) where "x |\| S \ fset_member S x" abbreviation - notin_fset :: "'a \ 'a fset \ bool" (infix "|\|" 50) + notin_fset :: "'a \ 'a fset \ bool" (infix \|\|\ 50) where "x |\| S \ \ (x |\| S)" @@ -1048,7 +1048,7 @@ and a proof of equivalence\ inductive - list_eq2 :: "'a list \ 'a list \ bool" ("_ \2 _") + list_eq2 :: "'a list \ 'a list \ bool" (\_ \2 _\) where "(a # b # xs) \2 (b # a # xs)" | "[] \2 []" @@ -1156,7 +1156,7 @@ \ no_notation - list_eq (infix "\" 50) and - list_eq2 (infix "\2" 50) + list_eq (infix \\\ 50) and + list_eq2 (infix \\2\ 50) end diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Quotient_Examples/Quotient_Int.thy --- a/src/HOL/Quotient_Examples/Quotient_Int.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Quotient_Examples/Quotient_Int.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ begin fun - intrel :: "(nat \ nat) \ (nat \ nat) \ bool" (infix "\" 50) + intrel :: "(nat \ nat) \ (nat \ nat) \ bool" (infix \\\ 50) where "intrel (x, y) (u, v) = (x + v = u + y)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Quotient_Examples/Quotient_Message.thy --- a/src/HOL/Quotient_Examples/Quotient_Message.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Quotient_Examples/Quotient_Message.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,7 +17,7 @@ | DECRYPT nat freemsg inductive - msgrel::"freemsg \ freemsg \ bool" (infixl "\" 50) + msgrel::"freemsg \ freemsg \ bool" (infixl \\\ 50) where CD: "CRYPT K (DECRYPT K X) \ X" | DC: "DECRYPT K (CRYPT K X) \ X" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Quotient_Examples/Quotient_Rat.thy --- a/src/HOL/Quotient_Examples/Quotient_Rat.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Quotient_Examples/Quotient_Rat.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,7 +9,7 @@ begin definition - ratrel :: "(int \ int) \ (int \ int) \ bool" (infix "\" 50) + ratrel :: "(int \ int) \ (int \ int) \ bool" (infix \\\ 50) where [simp]: "x \ y \ snd x \ 0 \ snd y \ 0 \ fst x * snd y = fst y * snd x" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Real_Asymp/Multiseries_Expansion.thy --- a/src/HOL/Real_Asymp/Multiseries_Expansion.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Real_Asymp/Multiseries_Expansion.thy Fri Sep 20 19:51:08 2024 +0200 @@ -2802,7 +2802,7 @@ inductive expands_to :: "(real \ real) \ 'a :: multiseries \ basis \ bool" - (infix "(expands'_to)" 50) where + (infix \(expands'_to)\ 50) where "is_expansion F basis \ eventually (\x. eval F x = f x) at_top \ (f expands_to F) basis" lemma dominant_term_expands_to: @@ -5375,7 +5375,7 @@ definition expansion_with_remainder_term :: "(real \ real) \ (real \ real) set \ bool" where "expansion_with_remainder_term _ _ = True" -notation (output) expansion_with_remainder_term (infixl "+" 10) +notation (output) expansion_with_remainder_term (infixl \+\ 10) ML_file \asymptotic_basis.ML\ ML_file \exp_log_expression.ML\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/SET_Protocol/Message_SET.thy --- a/src/HOL/SET_Protocol/Message_SET.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/SET_Protocol/Message_SET.thy Fri Sep 20 19:51:08 2024 +0200 @@ -70,9 +70,9 @@ (*Concrete syntax: messages appear as \\A,B,NA\\, etc...*) nonterminal mtuple_args syntax - "" :: "'a \ mtuple_args" ("_") - "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" ("_,/ _") - "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" ("(2\_,/ _\)") + "" :: "'a \ mtuple_args" (\_\) + "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" (\_,/ _\) + "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" (\(2\_,/ _\)\) syntax_consts "_MTuple_args" "_MTuple" \ MPair translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/SMT_Examples/Boogie.thy --- a/src/HOL/SMT_Examples/Boogie.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/SMT_Examples/Boogie.thy Fri Sep 20 19:51:08 2024 +0200 @@ -38,8 +38,8 @@ \ axiomatization - boogie_div :: "int \ int \ int" (infixl "div'_b" 70) and - boogie_mod :: "int \ int \ int" (infixl "mod'_b" 70) + boogie_div :: "int \ int \ int" (infixl \div'_b\ 70) and + boogie_mod :: "int \ int \ int" (infixl \mod'_b\ 70) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/SMT_Examples/SMT_Examples_Verit.thy --- a/src/HOL/SMT_Examples/SMT_Examples_Verit.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/SMT_Examples/SMT_Examples_Verit.thy Fri Sep 20 19:51:08 2024 +0200 @@ -683,8 +683,8 @@ fixes piecewise_C1 :: "('real :: {one,zero,ord} \ 'a :: {one,zero,ord}) \ 'real set \ bool" and joinpaths :: "('real \ 'a) \ ('real \ 'a) \ 'real \ 'a" begin -notation piecewise_C1 (infixr "piecewise'_C1'_differentiable'_on" 50) -notation joinpaths (infixr "+++" 75) +notation piecewise_C1 (infixr \piecewise'_C1'_differentiable'_on\ 50) +notation joinpaths (infixr \+++\ 75) lemma \(\v1. \v0. (rec_join v0 = v1 \ @@ -823,7 +823,7 @@ ground_resolution :: "'a \ 'a \ 'a \ bool" and is_least_false_clause :: "'afset \ 'a \ bool" and fset :: "'afset \ 'a set" and - union_fset :: "'afset \ 'afset \ 'afset" (infixr "|\|" 50) + union_fset :: "'afset \ 'afset \ 'afset" (infixr \|\|\ 50) begin term "a |\| b" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/SPARK/Manual/Reference.thy --- a/src/HOL/SPARK/Manual/Reference.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/SPARK/Manual/Reference.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,7 +9,7 @@ by (simp flip: mask_eq_exp_minus_1 take_bit_eq_mask take_bit_eq_mod) syntax (my_constrain output) - "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3) + "_constrain" :: "logic => type => logic" (\_ :: _\ [4, 0] 3) (*>*) chapter \HOL-\SPARK{} Reference\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/SPARK/SPARK_Setup.thy --- a/src/HOL/SPARK/SPARK_Setup.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/SPARK/SPARK_Setup.thy Fri Sep 20 19:51:08 2024 +0200 @@ -21,7 +21,7 @@ SPARK version of div, see section 4.4.1.1 of SPARK Proof Manual \ -definition sdiv :: "int \ int \ int" (infixl "sdiv" 70) where +definition sdiv :: "int \ int \ int" (infixl \sdiv\ 70) where "a sdiv b = sgn a * sgn b * (\a\ div \b\)" lemma sdiv_minus_dividend: "- a sdiv b = - (a sdiv b)" @@ -55,7 +55,7 @@ "fun_upds f xs y z = (if z \ xs then y else f z)" syntax - "_updsbind" :: "['a, 'a] => updbind" ("(2_ [:=]/ _)") + "_updsbind" :: "['a, 'a] => updbind" (\(2_ [:=]/ _)\) syntax_consts "_updsbind" == fun_upds diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Statespace/StateSpaceEx.thy --- a/src/HOL/Statespace/StateSpaceEx.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Statespace/StateSpaceEx.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,7 +10,7 @@ (*<*) syntax - "_statespace_updates" :: "('a \ 'b) \ updbinds \ ('a \ 'b)" ("_\_\" [900,0] 900) + "_statespace_updates" :: "('a \ 'b) \ updbinds \ ('a \ 'b)" (\_\_\\ [900,0] 900) (*>*) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Statespace/StateSpaceSyntax.thy --- a/src/HOL/Statespace/StateSpaceSyntax.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Statespace/StateSpaceSyntax.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,9 +11,9 @@ can choose if you want to use it or not.\ syntax - "_statespace_lookup" :: "('a \ 'b) \ 'a \ 'c" ("_\_" [60, 60] 60) + "_statespace_lookup" :: "('a \ 'b) \ 'a \ 'c" (\_\_\ [60, 60] 60) "_statespace_update" :: "('a \ 'b) \ 'a \ 'c \ ('a \ 'b)" - "_statespace_updates" :: "('a \ 'b) \ updbinds \ ('a \ 'b)" ("_<_>" [900, 0] 900) + "_statespace_updates" :: "('a \ 'b) \ updbinds \ ('a \ 'b)" (\_<_>\ [900, 0] 900) translations "_statespace_updates f (_updbinds b bs)" == diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/TLA/Action.thy --- a/src/HOL/TLA/Action.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/TLA/Action.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,16 +25,16 @@ syntax (* Syntax for writing action expressions in arbitrary contexts *) - "_ACT" :: "lift \ 'a" ("(ACT _)") + "_ACT" :: "lift \ 'a" (\(ACT _)\) - "_before" :: "lift \ lift" ("($_)" [100] 99) - "_after" :: "lift \ lift" ("(_$)" [100] 99) - "_unchanged" :: "lift \ lift" ("(unchanged _)" [100] 99) + "_before" :: "lift \ lift" (\($_)\ [100] 99) + "_after" :: "lift \ lift" (\(_$)\ [100] 99) + "_unchanged" :: "lift \ lift" (\(unchanged _)\ [100] 99) (*** Priming: same as "after" ***) - "_prime" :: "lift \ lift" ("(_`)" [100] 99) + "_prime" :: "lift \ lift" (\(_`)\ [100] 99) - "_Enabled" :: "lift \ lift" ("(Enabled _)" [100] 100) + "_Enabled" :: "lift \ lift" (\(Enabled _)\ [100] 100) translations "ACT A" => "(A::state*state \ _)" @@ -59,8 +59,8 @@ where angle_def: "AnAct A v \ ACT (A \ \ unchanged v)" syntax - "_SqAct" :: "[lift, lift] \ lift" ("([_]'_(_))" [0, 1000] 99) - "_AnAct" :: "[lift, lift] \ lift" ("(<_>'_(_))" [0, 1000] 99) + "_SqAct" :: "[lift, lift] \ lift" (\([_]'_(_))\ [0, 1000] 99) + "_AnAct" :: "[lift, lift] \ lift" (\(<_>'_(_))\ [0, 1000] 99) translations "_SqAct" == "CONST SqAct" "_AnAct" == "CONST AnAct" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/TLA/Init.thy --- a/src/HOL/TLA/Init.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/TLA/Init.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,8 +25,8 @@ where Init_def: "Initial F sigma = F (first_world sigma)" syntax - "_TEMP" :: "lift \ 'a" ("(TEMP _)") - "_Init" :: "lift \ lift" ("(Init _)"[40] 50) + "_TEMP" :: "lift \ 'a" (\(TEMP _)\) + "_Init" :: "lift \ lift" (\(Init _)\[40] 50) translations "TEMP F" => "(F::behavior \ _)" "_Init" == "CONST Initial" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/TLA/Intensional.thy --- a/src/HOL/TLA/Intensional.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/TLA/Intensional.thy Fri Sep 20 19:51:08 2024 +0200 @@ -33,11 +33,11 @@ where unl_lift3: "lift3 f x y z w \ f (x w) (y w) (z w)" (* "Rigid" quantification (logic level) *) -definition RAll :: "('a \ ('w::world) form) \ 'w form" (binder "Rall " 10) +definition RAll :: "('a \ ('w::world) form) \ 'w form" (binder \Rall \ 10) where unl_Rall: "(Rall x. A x) w \ \x. A x w" -definition REx :: "('a \ ('w::world) form) \ 'w form" (binder "Rex " 10) +definition REx :: "('a \ ('w::world) form) \ 'w form" (binder \Rex \ 10) where unl_Rex: "(Rex x. A x) w \ \x. A x w" -definition REx1 :: "('a \ ('w::world) form) \ 'w form" (binder "Rex! " 10) +definition REx1 :: "('a \ ('w::world) form) \ 'w form" (binder \Rex! \ 10) where unl_Rex1: "(Rex! x. A x) w \ \!x. A x w" @@ -46,56 +46,56 @@ nonterminal lift and liftargs syntax - "" :: "id \ lift" ("_") - "" :: "longid \ lift" ("_") - "" :: "var \ lift" ("_") - "_applC" :: "[lift, cargs] \ lift" ("(1_/ _)" [1000, 1000] 999) - "" :: "lift \ lift" ("'(_')") - "_lambda" :: "[idts, 'a] \ lift" ("(3\_./ _)" [0, 3] 3) - "_constrain" :: "[lift, type] \ lift" ("(_::_)" [4, 0] 3) - "" :: "lift \ liftargs" ("_") - "_liftargs" :: "[lift, liftargs] \ liftargs" ("_,/ _") - "_Valid" :: "lift \ bool" ("(\ _)" 5) - "_holdsAt" :: "['a, lift] \ bool" ("(_ \ _)" [100,10] 10) + "" :: "id \ lift" (\_\) + "" :: "longid \ lift" (\_\) + "" :: "var \ lift" (\_\) + "_applC" :: "[lift, cargs] \ lift" (\(1_/ _)\ [1000, 1000] 999) + "" :: "lift \ lift" (\'(_')\) + "_lambda" :: "[idts, 'a] \ lift" (\(3\_./ _)\ [0, 3] 3) + "_constrain" :: "[lift, type] \ lift" (\(_::_)\ [4, 0] 3) + "" :: "lift \ liftargs" (\_\) + "_liftargs" :: "[lift, liftargs] \ liftargs" (\_,/ _\) + "_Valid" :: "lift \ bool" (\(\ _)\ 5) + "_holdsAt" :: "['a, lift] \ bool" (\(_ \ _)\ [100,10] 10) (* Syntax for lifted expressions outside the scope of \ or |= *) - "_LIFT" :: "lift \ 'a" ("LIFT _") + "_LIFT" :: "lift \ 'a" (\LIFT _\) (* generic syntax for lifted constants and functions *) - "_const" :: "'a \ lift" ("(#_)" [1000] 999) - "_lift" :: "['a, lift] \ lift" ("(_<_>)" [1000] 999) - "_lift2" :: "['a, lift, lift] \ lift" ("(_<_,/ _>)" [1000] 999) - "_lift3" :: "['a, lift, lift, lift] \ lift" ("(_<_,/ _,/ _>)" [1000] 999) + "_const" :: "'a \ lift" (\(#_)\ [1000] 999) + "_lift" :: "['a, lift] \ lift" (\(_<_>)\ [1000] 999) + "_lift2" :: "['a, lift, lift] \ lift" (\(_<_,/ _>)\ [1000] 999) + "_lift3" :: "['a, lift, lift, lift] \ lift" (\(_<_,/ _,/ _>)\ [1000] 999) (* concrete syntax for common infix functions: reuse same symbol *) - "_liftEqu" :: "[lift, lift] \ lift" ("(_ =/ _)" [50,51] 50) - "_liftNeq" :: "[lift, lift] \ lift" ("(_ \/ _)" [50,51] 50) - "_liftNot" :: "lift \ lift" ("(\ _)" [40] 40) - "_liftAnd" :: "[lift, lift] \ lift" ("(_ \/ _)" [36,35] 35) - "_liftOr" :: "[lift, lift] \ lift" ("(_ \/ _)" [31,30] 30) - "_liftImp" :: "[lift, lift] \ lift" ("(_ \/ _)" [26,25] 25) - "_liftIf" :: "[lift, lift, lift] \ lift" ("(if (_)/ then (_)/ else (_))" 10) - "_liftPlus" :: "[lift, lift] \ lift" ("(_ +/ _)" [66,65] 65) - "_liftMinus" :: "[lift, lift] \ lift" ("(_ -/ _)" [66,65] 65) - "_liftTimes" :: "[lift, lift] \ lift" ("(_ */ _)" [71,70] 70) - "_liftDiv" :: "[lift, lift] \ lift" ("(_ div _)" [71,70] 70) - "_liftMod" :: "[lift, lift] \ lift" ("(_ mod _)" [71,70] 70) - "_liftLess" :: "[lift, lift] \ lift" ("(_/ < _)" [50, 51] 50) - "_liftLeq" :: "[lift, lift] \ lift" ("(_/ \ _)" [50, 51] 50) - "_liftMem" :: "[lift, lift] \ lift" ("(_/ \ _)" [50, 51] 50) - "_liftNotMem" :: "[lift, lift] \ lift" ("(_/ \ _)" [50, 51] 50) - "_liftFinset" :: "liftargs \ lift" ("{(_)}") + "_liftEqu" :: "[lift, lift] \ lift" (\(_ =/ _)\ [50,51] 50) + "_liftNeq" :: "[lift, lift] \ lift" (\(_ \/ _)\ [50,51] 50) + "_liftNot" :: "lift \ lift" (\(\ _)\ [40] 40) + "_liftAnd" :: "[lift, lift] \ lift" (\(_ \/ _)\ [36,35] 35) + "_liftOr" :: "[lift, lift] \ lift" (\(_ \/ _)\ [31,30] 30) + "_liftImp" :: "[lift, lift] \ lift" (\(_ \/ _)\ [26,25] 25) + "_liftIf" :: "[lift, lift, lift] \ lift" (\(if (_)/ then (_)/ else (_))\ 10) + "_liftPlus" :: "[lift, lift] \ lift" (\(_ +/ _)\ [66,65] 65) + "_liftMinus" :: "[lift, lift] \ lift" (\(_ -/ _)\ [66,65] 65) + "_liftTimes" :: "[lift, lift] \ lift" (\(_ */ _)\ [71,70] 70) + "_liftDiv" :: "[lift, lift] \ lift" (\(_ div _)\ [71,70] 70) + "_liftMod" :: "[lift, lift] \ lift" (\(_ mod _)\ [71,70] 70) + "_liftLess" :: "[lift, lift] \ lift" (\(_/ < _)\ [50, 51] 50) + "_liftLeq" :: "[lift, lift] \ lift" (\(_/ \ _)\ [50, 51] 50) + "_liftMem" :: "[lift, lift] \ lift" (\(_/ \ _)\ [50, 51] 50) + "_liftNotMem" :: "[lift, lift] \ lift" (\(_/ \ _)\ [50, 51] 50) + "_liftFinset" :: "liftargs \ lift" (\{(_)}\) (** TODO: syntax for lifted collection / comprehension **) - "_liftPair" :: "[lift,liftargs] \ lift" ("(1'(_,/ _'))") + "_liftPair" :: "[lift,liftargs] \ lift" (\(1'(_,/ _'))\) (* infix syntax for list operations *) - "_liftCons" :: "[lift, lift] \ lift" ("(_ #/ _)" [65,66] 65) - "_liftApp" :: "[lift, lift] \ lift" ("(_ @/ _)" [65,66] 65) - "_liftList" :: "liftargs \ lift" ("[(_)]") + "_liftCons" :: "[lift, lift] \ lift" (\(_ #/ _)\ [65,66] 65) + "_liftApp" :: "[lift, lift] \ lift" (\(_ @/ _)\ [65,66] 65) + "_liftList" :: "liftargs \ lift" (\[(_)]\) (* Rigid quantification (syntax level) *) - "_RAll" :: "[idts, lift] \ lift" ("(3\_./ _)" [0, 10] 10) - "_REx" :: "[idts, lift] \ lift" ("(3\_./ _)" [0, 10] 10) - "_REx1" :: "[idts, lift] \ lift" ("(3\!_./ _)" [0, 10] 10) + "_RAll" :: "[idts, lift] \ lift" (\(3\_./ _)\ [0, 10] 10) + "_REx" :: "[idts, lift] \ lift" (\(3\_./ _)\ [0, 10] 10) + "_REx1" :: "[idts, lift] \ lift" (\(3\!_./ _)\ [0, 10] 10) translations "_const" == "CONST const" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/TLA/Memory/ProcedureInterface.thy --- a/src/HOL/TLA/Memory/ProcedureInterface.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/TLA/Memory/ProcedureInterface.thy Fri Sep 20 19:51:08 2024 +0200 @@ -41,7 +41,7 @@ where "slice x i s \ x s i" syntax - "_slice" :: "[lift, 'a] \ lift" ("(_!_)" [70,70] 70) + "_slice" :: "[lift, 'a] \ lift" (\(_!_)\ [70,70] 70) translations "_slice" \ "CONST slice" @@ -67,8 +67,8 @@ \ (res$ = $v)" syntax - "_Call" :: "['a, 'b, lift] \ lift" ("(Call _ _ _)" [90,90,90] 90) - "_Return" :: "['a, 'b, lift] \ lift" ("(Return _ _ _)" [90,90,90] 90) + "_Call" :: "['a, 'b, lift] \ lift" (\(Call _ _ _)\ [90,90,90] 90) + "_Return" :: "['a, 'b, lift] \ lift" (\(Return _ _ _)\ [90,90,90] 90) translations "_Call" \ "CONST ACall" "_Return" \ "CONST AReturn" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/TLA/Stfun.thy --- a/src/HOL/TLA/Stfun.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/TLA/Stfun.thy Fri Sep 20 19:51:08 2024 +0200 @@ -33,8 +33,8 @@ stvars :: "'a stfun \ bool" syntax - "_PRED" :: "lift \ 'a" ("PRED _") - "_stvars" :: "lift \ bool" ("basevars _") + "_PRED" :: "lift \ 'a" (\PRED _\) + "_stvars" :: "lift \ bool" (\basevars _\) translations "PRED P" => "(P::state \ _)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/TLA/TLA.thy --- a/src/HOL/TLA/TLA.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/TLA/TLA.thy Fri Sep 20 19:51:08 2024 +0200 @@ -19,20 +19,20 @@ SF :: "[action, 'a stfun] \ temporal" (* Quantification over (flexible) state variables *) - EEx :: "('a stfun \ temporal) \ temporal" (binder "Eex " 10) - AAll :: "('a stfun \ temporal) \ temporal" (binder "Aall " 10) + EEx :: "('a stfun \ temporal) \ temporal" (binder \Eex \ 10) + AAll :: "('a stfun \ temporal) \ temporal" (binder \Aall \ 10) (** concrete syntax **) syntax - "_Box" :: "lift \ lift" ("(\_)" [40] 40) - "_Dmd" :: "lift \ lift" ("(\_)" [40] 40) - "_leadsto" :: "[lift,lift] \ lift" ("(_ \ _)" [23,22] 22) - "_stable" :: "lift \ lift" ("(stable/ _)") - "_WF" :: "[lift,lift] \ lift" ("(WF'(_')'_(_))" [0,60] 55) - "_SF" :: "[lift,lift] \ lift" ("(SF'(_')'_(_))" [0,60] 55) + "_Box" :: "lift \ lift" (\(\_)\ [40] 40) + "_Dmd" :: "lift \ lift" (\(\_)\ [40] 40) + "_leadsto" :: "[lift,lift] \ lift" (\(_ \ _)\ [23,22] 22) + "_stable" :: "lift \ lift" (\(stable/ _)\) + "_WF" :: "[lift,lift] \ lift" (\(WF'(_')'_(_))\ [0,60] 55) + "_SF" :: "[lift,lift] \ lift" (\(SF'(_')'_(_))\ [0,60] 55) - "_EEx" :: "[idts, lift] \ lift" ("(3\\ _./ _)" [0,10] 10) - "_AAll" :: "[idts, lift] \ lift" ("(3\\ _./ _)" [0,10] 10) + "_EEx" :: "[idts, lift] \ lift" (\(3\\ _./ _)\ [0,10] 10) + "_AAll" :: "[idts, lift] \ lift" (\(3\\ _./ _)\ [0,10] 10) translations "_Box" == "CONST Box" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Types_To_Sets/Examples/Linear_Algebra_On.thy --- a/src/HOL/Types_To_Sets/Examples/Linear_Algebra_On.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Types_To_Sets/Examples/Linear_Algebra_On.thy Fri Sep 20 19:51:08 2024 +0200 @@ -37,7 +37,7 @@ subsection \Definitions \<^emph>\on\ carrier set\ locale module_on = - fixes S and scale :: "'a::comm_ring_1 \ 'b::ab_group_add \ 'b" (infixr "*s" 75) + fixes S and scale :: "'a::comm_ring_1 \ 'b::ab_group_add \ 'b" (infixr \*s\ 75) assumes scale_right_distrib_on [algebra_simps]: "x \ S \ y \ S \ a *s (x + y) = a *s x + a *s y" and scale_left_distrib_on [algebra_simps]: "x \ S \ (a + b) *s x = a *s x + b *s x" and scale_scale_on [simp]: "x \ S \ a *s (b *s x) = (a * b) *s x" @@ -96,8 +96,8 @@ locale module_hom_on = m1: module_on S1 s1 + m2: module_on S2 s2 for S1 :: "'b::ab_group_add set" and S2 :: "'c::ab_group_add set" - and s1 :: "'a::comm_ring_1 \ 'b \ 'b" (infixr "*a" 75) - and s2 :: "'a::comm_ring_1 \ 'c \ 'c" (infixr "*b" 75) + + and s1 :: "'a::comm_ring_1 \ 'b \ 'b" (infixr \*a\ 75) + and s2 :: "'a::comm_ring_1 \ 'c \ 'c" (infixr \*b\ 75) + fixes f :: "'b \ 'c" assumes add: "\b1 b2. b1 \ S1 \ b2 \ S1 \ f (b1 + b2) = f b1 + f b2" and scale: "\b. b \ S1 \ f (r *a b) = r *b f b" @@ -109,7 +109,7 @@ by (auto intro!: ext) locale vector_space_on = module_on S scale - for S and scale :: "'a::field \ 'b::ab_group_add \ 'b" (infixr "*s" 75) + for S and scale :: "'a::field \ 'b::ab_group_add \ 'b" (infixr \*s\ 75) begin definition dim :: "'b set \ nat" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/Comp.thy --- a/src/HOL/UNITY/Comp.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/Comp.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,10 +25,10 @@ end -definition component_of :: "'a program =>'a program=> bool" (infixl "component'_of" 50) +definition component_of :: "'a program =>'a program=> bool" (infixl \component'_of\ 50) where "F component_of H == \G. F ok G & F\G = H" -definition strict_component_of :: "'a program\'a program=> bool" (infixl "strict'_component'_of" 50) +definition strict_component_of :: "'a program\'a program=> bool" (infixl \strict'_component'_of\ 50) where "F strict_component_of H == F component_of H & F\H" definition preserves :: "('a=>'b) => 'a program set" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/Constrains.thy --- a/src/HOL/UNITY/Constrains.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/Constrains.thy Fri Sep 20 19:51:08 2024 +0200 @@ -32,10 +32,10 @@ | Acts: "[| act \ Acts F; s \ reachable F; (s,s') \ act |] ==> s' \ reachable F" -definition Constrains :: "['a set, 'a set] => 'a program set" (infixl "Co" 60) where +definition Constrains :: "['a set, 'a set] => 'a program set" (infixl \Co\ 60) where "A Co B == {F. F \ (reachable F \ A) co B}" -definition Unless :: "['a set, 'a set] => 'a program set" (infixl "Unless" 60) where +definition Unless :: "['a set, 'a set] => 'a program set" (infixl \Unless\ 60) where "A Unless B == (A-B) Co (A \ B)" definition Stable :: "'a set => 'a program set" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/Detects.thy --- a/src/HOL/UNITY/Detects.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/Detects.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,10 +9,10 @@ theory Detects imports FP SubstAx begin -definition Detects :: "['a set, 'a set] => 'a program set" (infixl "Detects" 60) +definition Detects :: "['a set, 'a set] => 'a program set" (infixl \Detects\ 60) where "A Detects B = (Always (-A \ B)) \ (B LeadsTo A)" -definition Equality :: "['a set, 'a set] => 'a set" (infixl "<==>" 60) +definition Equality :: "['a set, 'a set] => 'a set" (infixl \<==>\ 60) where "A <==> B = (-A \ B) \ (A \ -B)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/ELT.thy --- a/src/HOL/UNITY/ELT.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/ELT.thy Fri Sep 20 19:51:08 2024 +0200 @@ -46,12 +46,12 @@ definition (*visible version of the LEADS-TO relation*) leadsETo :: "['a set, 'a set set, 'a set] => 'a program set" - ("(3_/ leadsTo[_]/ _)" [80,0,80] 80) + (\(3_/ leadsTo[_]/ _)\ [80,0,80] 80) where "leadsETo A CC B = {F. (A,B) \ elt CC F}" definition LeadsETo :: "['a set, 'a set set, 'a set] => 'a program set" - ("(3_/ LeadsTo[_]/ _)" [80,0,80] 80) + (\(3_/ LeadsTo[_]/ _)\ [80,0,80] 80) where "LeadsETo A CC B = {F. F \ (reachable F Int A) leadsTo[(%C. reachable F Int C) ` CC] B}" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/Follows.thy --- a/src/HOL/UNITY/Follows.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/Follows.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,7 +9,7 @@ imports SubstAx ListOrder "HOL-Library.Multiset" begin -definition Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set" (infixl "Fols" 65) where +definition Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set" (infixl \Fols\ 65) where "f Fols g == Increasing g \ Increasing f Int Always {s. f s \ g s} Int (\k. {s. k \ g s} LeadsTo {s. k \ f s})" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/Guar.thy --- a/src/HOL/UNITY/Guar.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/Guar.thy Fri Sep 20 19:51:08 2024 +0200 @@ -38,7 +38,7 @@ text\Guarantees properties\ -definition guar :: "['a program set, 'a program set] => 'a program set" (infixl "guarantees" 55) where +definition guar :: "['a program set, 'a program set] => 'a program set" (infixl \guarantees\ 55) where (*higher than membership, lower than Co*) "X guarantees Y == {F. \G. F ok G --> F\G \ X --> F\G \ Y}" @@ -56,13 +56,13 @@ "welldef == {F. Init F \ {}}" definition refines :: "['a program, 'a program, 'a program set] => bool" - ("(3_ refines _ wrt _)" [10,10,10] 10) where + (\(3_ refines _ wrt _)\ [10,10,10] 10) where "G refines F wrt X == \H. (F ok H & G ok H & F\H \ welldef \ X) --> (G\H \ welldef \ X)" definition iso_refines :: "['a program, 'a program, 'a program set] => bool" - ("(3_ iso'_refines _ wrt _)" [10,10,10] 10) where + (\(3_ iso'_refines _ wrt _)\ [10,10,10] 10) where "G iso_refines F wrt X == F \ welldef \ X --> G \ welldef \ X" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/ListOrder.thy --- a/src/HOL/UNITY/ListOrder.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/ListOrder.thy Fri Sep 20 19:51:08 2024 +0200 @@ -49,11 +49,11 @@ "Ge == {(x,y). y <= x}" abbreviation - pfixLe :: "[nat list, nat list] => bool" (infixl "pfixLe" 50) where + pfixLe :: "[nat list, nat list] => bool" (infixl \pfixLe\ 50) where "xs pfixLe ys == (xs,ys) \ genPrefix Le" abbreviation - pfixGe :: "[nat list, nat list] => bool" (infixl "pfixGe" 50) where + pfixGe :: "[nat list, nat list] => bool" (infixl \pfixGe\ 50) where "xs pfixGe ys == (xs,ys) \ genPrefix Ge" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/PPROD.thy --- a/src/HOL/UNITY/PPROD.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/PPROD.thy Fri Sep 20 19:51:08 2024 +0200 @@ -15,7 +15,7 @@ "PLam I F == \i \ I. lift i (F i)" syntax - "_PLam" :: "[pttrn, nat set, 'b set] => (nat => 'b) set" ("(3plam _:_./ _)" 10) + "_PLam" :: "[pttrn, nat set, 'b set] => (nat => 'b) set" (\(3plam _:_./ _)\ 10) syntax_consts "_PLam" == PLam translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/SubstAx.thy --- a/src/HOL/UNITY/SubstAx.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/SubstAx.thy Fri Sep 20 19:51:08 2024 +0200 @@ -9,13 +9,13 @@ theory SubstAx imports WFair Constrains begin -definition Ensures :: "['a set, 'a set] => 'a program set" (infixl "Ensures" 60) where +definition Ensures :: "['a set, 'a set] => 'a program set" (infixl \Ensures\ 60) where "A Ensures B == {F. F \ (reachable F \ A) ensures B}" -definition LeadsTo :: "['a set, 'a set] => 'a program set" (infixl "LeadsTo" 60) where +definition LeadsTo :: "['a set, 'a set] => 'a program set" (infixl \LeadsTo\ 60) where "A LeadsTo B == {F. F \ (reachable F \ A) leadsTo B}" -notation LeadsTo (infixl "\w" 60) +notation LeadsTo (infixl \\w\ 60) text\Resembles the previous definition of LeadsTo\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/UNITY.thy --- a/src/HOL/UNITY/UNITY.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/UNITY.thy Fri Sep 20 19:51:08 2024 +0200 @@ -24,10 +24,10 @@ definition Acts :: "'a program => ('a * 'a)set set" where "Acts F == (%(init, acts, allowed). acts) (Rep_Program F)" -definition "constrains" :: "['a set, 'a set] => 'a program set" (infixl "co" 60) where +definition "constrains" :: "['a set, 'a set] => 'a program set" (infixl \co\ 60) where "A co B == {F. \act \ Acts F. act``A \ B}" -definition unless :: "['a set, 'a set] => 'a program set" (infixl "unless" 60) where +definition unless :: "['a set, 'a set] => 'a program set" (infixl \unless\ 60) where "A unless B == (A-B) co (A \ B)" definition mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set) diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/Union.thy --- a/src/HOL/UNITY/Union.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/Union.thy Fri Sep 20 19:51:08 2024 +0200 @@ -11,7 +11,7 @@ (*FIXME: conjoin Init F \ Init G \ {} *) definition - ok :: "['a program, 'a program] => bool" (infixl "ok" 65) + ok :: "['a program, 'a program] => bool" (infixl \ok\ 65) where "F ok G == Acts F \ AllowedActs G & Acts G \ AllowedActs F" @@ -26,11 +26,11 @@ \i \ I. AllowedActs (F i))" definition - Join :: "['a program, 'a program] => 'a program" (infixl "\" 65) + Join :: "['a program, 'a program] => 'a program" (infixl \\\ 65) where "F \ G = mk_program (Init F \ Init G, Acts F \ Acts G, AllowedActs F \ AllowedActs G)" -definition SKIP :: "'a program" ("\") +definition SKIP :: "'a program" (\\\) where "\ = mk_program (UNIV, {}, UNIV)" (*Characterizes safety properties. Used with specifying Allowed*) @@ -39,8 +39,8 @@ where "safety_prop X \ SKIP \ X \ (\G. Acts G \ \(Acts ` X) \ G \ X)" syntax - "_JOIN1" :: "[pttrns, 'b set] => 'b set" ("(3\_./ _)" 10) - "_JOIN" :: "[pttrn, 'a set, 'b set] => 'b set" ("(3\_\_./ _)" 10) + "_JOIN1" :: "[pttrns, 'b set] => 'b set" (\(3\_./ _)\ 10) + "_JOIN" :: "[pttrn, 'a set, 'b set] => 'b set" (\(3\_\_./ _)\ 10) syntax_consts "_JOIN1" "_JOIN" == JOIN translations diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/UNITY/WFair.thy --- a/src/HOL/UNITY/WFair.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/UNITY/WFair.thy Fri Sep 20 19:51:08 2024 +0200 @@ -42,7 +42,7 @@ "transient A == {F. \act\Acts F. act``A \ -A}" definition - ensures :: "['a set, 'a set] => 'a program set" (infixl "ensures" 60) where + ensures :: "['a set, 'a set] => 'a program set" (infixl \ensures\ 60) where "A ensures B == (A-B co A \ B) \ transient (A-B)" @@ -59,7 +59,7 @@ | Union: "\A \ S. (A,B) \ leads F ==> (Union S, B) \ leads F" -definition leadsTo :: "['a set, 'a set] => 'a program set" (infixl "leadsTo" 60) where +definition leadsTo :: "['a set, 'a set] => 'a program set" (infixl \leadsTo\ 60) where \ \visible version of the LEADS-TO relation\ "A leadsTo B == {F. (A,B) \ leads F}" @@ -67,7 +67,7 @@ \ \predicate transformer: the largest set that leads to \<^term>\B\\ "wlt F B == \{A. F \ A leadsTo B}" -notation leadsTo (infixl "\" 60) +notation leadsTo (infixl \\\ 60) subsection\transient\ diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Unix/Unix.thy --- a/src/HOL/Unix/Unix.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Unix/Unix.thy Fri Sep 20 19:51:08 2024 +0200 @@ -336,7 +336,7 @@ \ inductive transition :: "file \ operation \ file \ bool" - ("_ \_\ _" [90, 1000, 90] 100) + (\_ \_\ _\ [90, 1000, 90] 100) where read: "root \(Read uid text path)\ root" @@ -475,7 +475,7 @@ running processes over a finite amount of time. \ -inductive transitions :: "file \ operation list \ file \ bool" ("_ \_\ _" [90, 1000, 90] 100) +inductive transitions :: "file \ operation list \ file \ bool" (\_ \_\ _\ [90, 1000, 90] 100) where nil: "root \[]\ root" | cons: "root \(x # xs)\ root''" if "root \x\ root'" and "root' \xs\ root''" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ZF/HOLZF.thy --- a/src/HOL/ZF/HOLZF.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ZF/HOLZF.thy Fri Sep 20 19:51:08 2024 +0200 @@ -188,7 +188,7 @@ definition Field :: "ZF \ ZF" where "Field A == union (Domain A) (Range A)" -definition app :: "ZF \ ZF => ZF" (infixl "\" 90) \ \function application\ where +definition app :: "ZF \ ZF => ZF" (infixl \\\ 90) \ \function application\ where "f \ x == (THE y. Elem (Opair x y) f)" definition isFun :: "ZF \ bool" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Antiquote.thy --- a/src/HOL/ex/Antiquote.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Antiquote.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,14 +10,14 @@ text \A simple example on quote / antiquote in higher-order abstract syntax.\ -definition var :: "'a \ ('a \ nat) \ nat" ("VAR _" [1000] 999) +definition var :: "'a \ ('a \ nat) \ nat" (\VAR _\ [1000] 999) where "var x env = env x" definition Expr :: "(('a \ nat) \ nat) \ ('a \ nat) \ nat" where "Expr exp env = exp env" syntax - "_Expr" :: "'a \ 'a" ("EXPR _" [1000] 999) + "_Expr" :: "'a \ 'a" (\EXPR _\ [1000] 999) parse_translation \[Syntax_Trans.quote_antiquote_tr diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/BigO.thy --- a/src/HOL/ex/BigO.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/BigO.thy Fri Sep 20 19:51:08 2024 +0200 @@ -35,7 +35,7 @@ subsection \Definitions\ -definition bigo :: "('a \ 'b::linordered_idom) \ ('a \ 'b) set" ("(1O'(_'))") +definition bigo :: "('a \ 'b::linordered_idom) \ ('a \ 'b) set" (\(1O'(_'))\) where "O(f:: 'a \ 'b) = {h. \c. \x. \h x\ \ c * \f x\}" lemma bigo_pos_const: @@ -407,7 +407,7 @@ subsection \Less than or equal to\ -definition lesso :: "('a \ 'b::linordered_idom) \ ('a \ 'b) \ 'a \ 'b" (infixl " 'b::linordered_idom) \ ('a \ 'b) \ 'a \ 'b" (infixl \ 70) where "f x. max (f x - g x) 0)" lemma bigo_lesseq1: "f =o O(h) \ \x. \g x\ \ \f x\ \ g =o O(h)" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/CTL.thy --- a/src/HOL/ex/CTL.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/CTL.thy Fri Sep 20 19:51:08 2024 +0200 @@ -24,7 +24,7 @@ type_synonym 'a ctl = "'a set" -definition imp :: "'a ctl \ 'a ctl \ 'a ctl" (infixr "\" 75) +definition imp :: "'a ctl \ 'a ctl \ 'a ctl" (infixr \\\ 75) where "p \ q = - p \ q" lemma [intro!]: "p \ p \ q \ q" unfolding imp_def by auto @@ -53,13 +53,13 @@ \<^cite>\"McMillan-PhDThesis"\. \ -definition EX ("\<^bold>E\<^bold>X _" [80] 90) +definition EX (\\<^bold>E\<^bold>X _\ [80] 90) where [simp]: "\<^bold>E\<^bold>X p = {s. \s'. (s, s') \ \ \ s' \ p}" -definition EF ("\<^bold>E\<^bold>F _" [80] 90) +definition EF (\\<^bold>E\<^bold>F _\ [80] 90) where [simp]: "\<^bold>E\<^bold>F p = lfp (\s. p \ \<^bold>E\<^bold>X s)" -definition EG ("\<^bold>E\<^bold>G _" [80] 90) +definition EG (\\<^bold>E\<^bold>G _\ [80] 90) where [simp]: "\<^bold>E\<^bold>G p = gfp (\s. p \ \<^bold>E\<^bold>X s)" text \ @@ -67,11 +67,11 @@ \\<^bold>E\<^bold>F\ and \\<^bold>E\<^bold>G\. \ -definition AX ("\<^bold>A\<^bold>X _" [80] 90) +definition AX (\\<^bold>A\<^bold>X _\ [80] 90) where [simp]: "\<^bold>A\<^bold>X p = - \<^bold>E\<^bold>X - p" -definition AF ("\<^bold>A\<^bold>F _" [80] 90) +definition AF (\\<^bold>A\<^bold>F _\ [80] 90) where [simp]: "\<^bold>A\<^bold>F p = - \<^bold>E\<^bold>G - p" -definition AG ("\<^bold>A\<^bold>G _" [80] 90) +definition AG (\\<^bold>A\<^bold>G _\ [80] 90) where [simp]: "\<^bold>A\<^bold>G p = - \<^bold>E\<^bold>F - p" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Cartouche_Examples.thy --- a/src/HOL/ex/Cartouche_Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Cartouche_Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -78,7 +78,7 @@ end; \ -syntax "_cartouche_string" :: \cartouche_position \ string\ ("_") +syntax "_cartouche_string" :: \cartouche_position \ string\ (\_\) parse_translation \ [(\<^syntax_const>\_cartouche_string\, @@ -95,7 +95,7 @@ subsubsection \Nested quotes\ -syntax "_string_string" :: \string_position \ string\ ("_") +syntax "_string_string" :: \string_position \ string\ (\_\) parse_translation \ [(\<^syntax_const>\_string_string\, K (string_tr Lexicon.explode_string))] diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Chinese.thy --- a/src/HOL/ex/Chinese.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Chinese.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,18 +17,18 @@ \ datatype shuzi = - One ("一") - | Two ("二") - | Three ("三") - | Four ("四") - | Five ("五") - | Six ("六") - | Seven ("七") - | Eight ("八") - | Nine ("九") - | Ten ("十") + One (\一\) + | Two (\二\) + | Three (\三\) + | Four (\四\) + | Five (\五\) + | Six (\六\) + | Seven (\七\) + | Eight (\八\) + | Nine (\九\) + | Ten (\十\) -primrec translate :: "shuzi \ nat" ("|_|" [100] 100) where +primrec translate :: "shuzi \ nat" (\|_|\ [100] 100) where "|一| = 1" |"|二| = 2" |"|三| = 3" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Code_Lazy_Demo.thy --- a/src/HOL/ex/Code_Lazy_Demo.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Code_Lazy_Demo.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ text \Lazy evaluation for streams\ codatatype 'a stream = - SCons (shd: 'a) (stl: "'a stream") (infixr "##" 65) + SCons (shd: 'a) (stl: "'a stream") (infixr \##\ 65) primcorec up :: "nat \ nat stream" where "up n = n ## up (n + 1)" @@ -36,14 +36,14 @@ text \Lazy types need not be infinite. We can also have lazy types that are finite.\ datatype 'a llist - = LNil ("\<^bold>\\<^bold>\") - | LCons (lhd: 'a) (ltl: "'a llist") (infixr "###" 65) + = LNil (\\<^bold>\\<^bold>\\) + | LCons (lhd: 'a) (ltl: "'a llist") (infixr \###\ 65) nonterminal llist_args syntax - "" :: "'a \ llist_args" ("_") - "_llist_args" :: "'a \ llist_args \ llist_args" ("_,/ _") - "_llist" :: "llist_args => 'a list" ("\<^bold>\(_)\<^bold>\") + "" :: "'a \ llist_args" (\_\) + "_llist_args" :: "'a \ llist_args \ llist_args" (\_,/ _\) + "_llist" :: "llist_args => 'a list" (\\<^bold>\(_)\<^bold>\\) syntax_consts "_llist_args" "_llist" == LCons translations @@ -81,7 +81,7 @@ A conversion function to lazy lists is enough.\ primrec lappend :: "'a llist \ 'a llist \ 'a llist" - (infixr "@@" 65) where + (infixr \@@\ 65) where "\<^bold>\\<^bold>\ @@ ys = ys" | "(x ### xs) @@ ys = x ### (xs @@ ys)" @@ -113,10 +113,10 @@ section \Branching datatypes\ datatype tree - = L ("\") - | Node tree tree (infix "\" 900) + = L (\\\) + | Node tree tree (infix \\\ 900) -notation (output) Node ("\(//\<^bold>l: _//\<^bold>r: _)") +notation (output) Node (\\(//\<^bold>l: _//\<^bold>r: _)\) code_lazy_type tree diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Function_Growth.thy --- a/src/HOL/ex/Function_Growth.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Function_Growth.thy Fri Sep 20 19:51:08 2024 +0200 @@ -40,7 +40,7 @@ subsection \The \\\ relation\ -definition less_eq_fun :: "(nat \ nat) \ (nat \ nat) \ bool" (infix "\" 50) +definition less_eq_fun :: "(nat \ nat) \ (nat \ nat) \ bool" (infix \\\ 50) where "f \ g \ (\c>0. \n. \m>n. f m \ c * g m)" @@ -73,7 +73,7 @@ subsection \The \\\ relation, the equivalence relation induced by \\\\ -definition equiv_fun :: "(nat \ nat) \ (nat \ nat) \ bool" (infix "\" 50) +definition equiv_fun :: "(nat \ nat) \ (nat \ nat) \ bool" (infix \\\ 50) where "f \ g \ (\c\<^sub>1>0. \c\<^sub>2>0. \n. \m>n. f m \ c\<^sub>1 * g m \ g m \ c\<^sub>2 * f m)" @@ -110,7 +110,7 @@ subsection \The \\\ relation, the strict part of \\\\ -definition less_fun :: "(nat \ nat) \ (nat \ nat) \ bool" (infix "\" 50) +definition less_fun :: "(nat \ nat) \ (nat \ nat) \ bool" (infix \\\ 50) where "f \ g \ f \ g \ \ g \ f" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Hebrew.thy --- a/src/HOL/ex/Hebrew.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Hebrew.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,28 +17,28 @@ subsection \The Hebrew Alef-Bet (א-ב).\ datatype alef_bet = - Alef ("א") - | Bet ("ב") - | Gimel ("ג") - | Dalet ("ד") - | He ("ה") - | Vav ("ו") - | Zayin ("ז") - | Het ("ח") - | Tet ("ט") - | Yod ("י") - | Kaf ("כ") - | Lamed ("ל") - | Mem ("מ") - | Nun ("נ") - | Samekh ("ס") - | Ayin ("ע") - | Pe ("פ") - | Tsadi ("צ") - | Qof ("ק") - | Resh ("ר") - | Shin ("ש") - | Tav ("ת") + Alef (\א\) + | Bet (\ב\) + | Gimel (\ג\) + | Dalet (\ד\) + | He (\ה\) + | Vav (\ו\) + | Zayin (\ז\) + | Het (\ח\) + | Tet (\ט\) + | Yod (\י\) + | Kaf (\כ\) + | Lamed (\ל\) + | Mem (\מ\) + | Nun (\נ\) + | Samekh (\ס\) + | Ayin (\ע\) + | Pe (\פ\) + | Tsadi (\צ\) + | Qof (\ק\) + | Resh (\ר\) + | Shin (\ש\) + | Tav (\ת\) thm alef_bet.induct diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/LocaleTest2.thy --- a/src/HOL/ex/LocaleTest2.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/LocaleTest2.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,7 +27,7 @@ subsubsection \Definitions\ locale dpo = - fixes le :: "['a, 'a] => bool" (infixl "\" 50) + fixes le :: "['a, 'a] => bool" (infixl \\\ 50) assumes refl [intro, simp]: "x \ x" and antisym [intro]: "[| x \ y; y \ x |] ==> x = y" and trans [trans]: "[| x \ y; y \ z |] ==> x \ z" @@ -39,7 +39,7 @@ by (blast intro: trans) definition - less :: "['a, 'a] => bool" (infixl "\" 50) + less :: "['a, 'a] => bool" (infixl \\\ 50) where "(x \ y) = (x \ y & x ~= y)" theorem abs_test: @@ -63,11 +63,11 @@ begin definition - meet :: "['a, 'a] => 'a" (infixl "\" 70) + meet :: "['a, 'a] => 'a" (infixl \\\ 70) where "x \ y = (THE inf. is_inf x y inf)" definition - join :: "['a, 'a] => 'a" (infixl "\" 65) + join :: "['a, 'a] => 'a" (infixl \\\ 65) where "x \ y = (THE sup. is_sup x y sup)" lemma is_infI [intro?]: "i \ x \ i \ y \ @@ -620,7 +620,7 @@ subsubsection \Locale declarations and lemmas\ locale Dsemi = - fixes prod (infixl "**" 65) + fixes prod (infixl \**\ 65) assumes assoc: "(x ** y) ** z = x ** (y ** z)" locale Dmonoid = Dsemi + @@ -791,7 +791,7 @@ locale Dhom = prod: Dgrp prod one + sum: Dgrp sum zero - for prod (infixl "**" 65) and one and sum (infixl "+++" 60) and zero + + for prod (infixl \**\ 65) and one and sum (infixl \+++\ 60) and zero + fixes hom assumes hom_mult [simp]: "hom (x ** y) = hom x +++ hom y" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Multiquote.thy --- a/src/HOL/ex/Multiquote.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Multiquote.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,8 +14,8 @@ \ syntax - "_quote" :: "'b \ ('a \ 'b)" ("\_\" [0] 1000) - "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) + "_quote" :: "'b \ ('a \ 'b)" (\\_\\ [0] 1000) + "_antiquote" :: "('a \ 'b) \ 'b" (\\_\ [1000] 1000) parse_translation \ let diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/PER.thy --- a/src/HOL/ex/PER.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/PER.thy Fri Sep 20 19:51:08 2024 +0200 @@ -32,7 +32,7 @@ \ class partial_equiv = - fixes eqv :: "'a \ 'a \ bool" (infixl "\" 50) + fixes eqv :: "'a \ 'a \ bool" (infixl \\\ 50) assumes partial_equiv_sym [elim?]: "x \ y \ y \ x" assumes partial_equiv_trans [trans]: "x \ y \ y \ z \ x \ z" @@ -168,7 +168,7 @@ representation of elements of a quotient type. \ -definition eqv_class :: "('a::partial_equiv) \ 'a quot" ("\_\") +definition eqv_class :: "('a::partial_equiv) \ 'a quot" (\\_\\) where "\a\ = Abs_quot {x. a \ x}" theorem quot_rep: "\a. A = \a\" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Serbian.thy --- a/src/HOL/ex/Serbian.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Serbian.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,68 +12,68 @@ text \Serbian cyrillic letters.\ datatype azbuka = - azbA ("А") -| azbB ("Б") -| azbV ("В") -| azbG ("Г") -| azbD ("Д") -| azbDj ("Ђ") -| azbE ("Е") -| azbZv ("Ж") -| azbZ ("З") -| azbI ("И") -| azbJ ("Ј") -| azbK ("К") -| azbL ("Л") -| azbLj ("Љ") -| azbM ("М") -| azbN ("Н") -| azbNj ("Њ") -| azbO ("О") -| azbP ("П") -| azbR ("Р") -| azbS ("С") -| azbT ("Т") -| azbC' ("Ћ") -| azbU ("У") -| azbF ("Ф") -| azbH ("Х") -| azbC ("Ц") -| azbCv ("Ч") -| azbDzv ("Џ") -| azbSv ("Ш") + azbA (\А\) +| azbB (\Б\) +| azbV (\В\) +| azbG (\Г\) +| azbD (\Д\) +| azbDj (\Ђ\) +| azbE (\Е\) +| azbZv (\Ж\) +| azbZ (\З\) +| azbI (\И\) +| azbJ (\Ј\) +| azbK (\К\) +| azbL (\Л\) +| azbLj (\Љ\) +| azbM (\М\) +| azbN (\Н\) +| azbNj (\Њ\) +| azbO (\О\) +| azbP (\П\) +| azbR (\Р\) +| azbS (\С\) +| azbT (\Т\) +| azbC' (\Ћ\) +| azbU (\У\) +| azbF (\Ф\) +| azbH (\Х\) +| azbC (\Ц\) +| azbCv (\Ч\) +| azbDzv (\Џ\) +| azbSv (\Ш\) | azbSpc thm azbuka.induct text \Serbian latin letters.\ datatype abeceda = - abcA ("A") -| abcB ("B") -| abcC ("C") -| abcCv ("Č") -| abcC' ("Ć") -| abcD ("D") -| abcE ("E") -| abcF ("F") -| abcG ("G") -| abcH ("H") -| abcI ("I") -| abcJ ("J") -| abcK ("K") -| abcL ("L") -| abcM ("M") -| abcN ("N") -| abcO ("O") -| abcP ("P") -| abcR ("R") -| abcS ("S") -| abcSv ("Š") -| abcT ("T") -| abcU ("U") -| abcV ("V") -| abcZ ("Z") -| abcvZ ("Ž") + abcA (\A\) +| abcB (\B\) +| abcC (\C\) +| abcCv (\Č\) +| abcC' (\Ć\) +| abcD (\D\) +| abcE (\E\) +| abcF (\F\) +| abcG (\G\) +| abcH (\H\) +| abcI (\I\) +| abcJ (\J\) +| abcK (\K\) +| abcL (\L\) +| abcM (\M\) +| abcN (\N\) +| abcO (\O\) +| abcP (\P\) +| abcR (\R\) +| abcS (\S\) +| abcSv (\Š\) +| abcT (\T\) +| abcU (\U\) +| abcV (\V\) +| abcZ (\Z\) +| abcvZ (\Ž\) | abcSpc thm abeceda.induct diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Sudoku.thy --- a/src/HOL/ex/Sudoku.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Sudoku.thy Fri Sep 20 19:51:08 2024 +0200 @@ -17,9 +17,9 @@ \nitpick\, which is used below.) \ -no_notation Groups.one_class.one ("1") +no_notation Groups.one_class.one (\1\) -datatype digit = A ("1") | B ("2") | C ("3") | D ("4") | E ("5") | F ("6") | G ("7") | H ("8") | I ("9") +datatype digit = A (\1\) | B (\2\) | C (\3\) | D (\4\) | E (\5\) | F (\6\) | G (\7\) | H (\8\) | I (\9\) definition valid :: "digit => digit => digit => digit => digit => digit => digit => digit => digit => bool" where diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Tarski.thy --- a/src/HOL/ex/Tarski.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Tarski.thy Fri Sep 20 19:51:08 2024 +0200 @@ -80,7 +80,7 @@ (SIGMA cl : CompleteLattice. {S. S \ pset cl \ \pset = S, order = induced S (order cl)\ \ CompleteLattice})" -abbreviation sublat :: "['a set, 'a potype] \ bool" ("_ <<= _" [51, 50] 50) +abbreviation sublat :: "['a set, 'a potype] \ bool" (\_ <<= _\ [51, 50] 50) where "S <<= cl \ S \ sublattice `` {cl}" definition dual :: "'a potype \ 'a potype" diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/ex/Unification.thy --- a/src/HOL/ex/Unification.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/ex/Unification.thy Fri Sep 20 19:51:08 2024 +0200 @@ -35,7 +35,7 @@ datatype 'a trm = Var 'a | Const 'a - | Comb "'a trm" "'a trm" (infix "\" 60) + | Comb "'a trm" "'a trm" (infix \\\ 60) primrec vars_of :: "'a trm \ 'a set" where @@ -43,7 +43,7 @@ | "vars_of (Const c) = {}" | "vars_of (M \ N) = vars_of M \ vars_of N" -fun occs :: "'a trm \ 'a trm \ bool" (infixl "\" 54) +fun occs :: "'a trm \ 'a trm \ bool" (infixl \\\ 54) where "u \ Var v \ False" | "u \ Const c \ False" @@ -77,17 +77,17 @@ "assoc x d [] = d" | "assoc x d ((p,q)#t) = (if x = p then q else assoc x d t)" -primrec subst :: "'a trm \ 'a subst \ 'a trm" (infixl "\" 55) +primrec subst :: "'a trm \ 'a subst \ 'a trm" (infixl \\\ 55) where "(Var v) \ s = assoc v (Var v) s" | "(Const c) \ s = (Const c)" | "(M \ N) \ s = (M \ s) \ (N \ s)" -definition subst_eq (infixr "\" 52) +definition subst_eq (infixr \\\ 52) where "s1 \ s2 \ (\t. t \ s1 = t \ s2)" -fun comp :: "'a subst \ 'a subst \ 'a subst" (infixl "\" 56) +fun comp :: "'a subst \ 'a subst \ 'a subst" (infixl \\\ 56) where "[] \ bl = bl" | "((a,b) # al) \ bl = (a, b \ bl) # (al \ bl)" diff -r 46f59511b7bb -r d97fdabd9e2b src/LCF/LCF.thy --- a/src/LCF/LCF.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/LCF/LCF.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,8 +18,8 @@ typedecl tr typedecl void -typedecl ('a,'b) prod (infixl "*" 6) -typedecl ('a,'b) sum (infixl "+" 5) +typedecl ('a,'b) prod (infixl \*\ 6) +typedecl ('a,'b) sum (infixl \+\ 5) instance "fun" :: (cpo, cpo) cpo .. instance prod :: (cpo, cpo) cpo .. @@ -38,10 +38,10 @@ INR :: "'b \ 'a+'b" WHEN :: "['a\'c, 'b\'c, 'a+'b] \ 'c" adm :: "('a \ o) \ o" - VOID :: "void" ("'(')") - PAIR :: "['a,'b] \ 'a*'b" ("(1<_,/_>)" [0,0] 100) - COND :: "[tr,'a,'a] \ 'a" ("(_ \/ (_ |/ _))" [60,60,60] 60) - less :: "['a,'a] \ o" (infixl "<<" 50) + VOID :: "void" (\'(')\) + PAIR :: "['a,'b] \ 'a*'b" (\(1<_,/_>)\ [0,0] 100) + COND :: "[tr,'a,'a] \ 'a" (\(_ \/ (_ |/ _))\ [60,60,60] 60) + less :: "['a,'a] \ o" (infixl \<<\ 50) axiomatization where (** DOMAIN THEORY **) diff -r 46f59511b7bb -r d97fdabd9e2b src/Pure/Examples/First_Order_Logic.thy --- a/src/Pure/Examples/First_Order_Logic.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Pure/Examples/First_Order_Logic.thy Fri Sep 20 19:51:08 2024 +0200 @@ -18,21 +18,21 @@ typedecl i typedecl o -judgment Trueprop :: "o \ prop" ("_" 5) +judgment Trueprop :: "o \ prop" (\_\ 5) subsection \Propositional logic\ -axiomatization false :: o ("\") +axiomatization false :: o (\\\) where falseE [elim]: "\ \ A" -axiomatization imp :: "o \ o \ o" (infixr "\" 25) +axiomatization imp :: "o \ o \ o" (infixr \\\ 25) where impI [intro]: "(A \ B) \ A \ B" and mp [dest]: "A \ B \ A \ B" -axiomatization conj :: "o \ o \ o" (infixr "\" 35) +axiomatization conj :: "o \ o \ o" (infixr \\\ 35) where conjI [intro]: "A \ B \ A \ B" and conjD1: "A \ B \ A" and conjD2: "A \ B \ B" @@ -48,20 +48,20 @@ qed -axiomatization disj :: "o \ o \ o" (infixr "\" 30) +axiomatization disj :: "o \ o \ o" (infixr \\\ 30) where disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" and disjI1 [intro]: "A \ A \ B" and disjI2 [intro]: "B \ A \ B" -definition true :: o ("\") +definition true :: o (\\\) where "\ \ \ \ \" theorem trueI [intro]: \ unfolding true_def .. -definition not :: "o \ o" ("\ _" [40] 40) +definition not :: "o \ o" (\\ _\ [40] 40) where "\ A \ A \ \" theorem notI [intro]: "(A \ \) \ \ A" @@ -76,7 +76,7 @@ qed -definition iff :: "o \ o \ o" (infixr "\" 25) +definition iff :: "o \ o \ o" (infixr \\\ 25) where "A \ B \ (A \ B) \ (B \ A)" theorem iffI [intro]: @@ -112,7 +112,7 @@ subsection \Equality\ -axiomatization equal :: "i \ i \ o" (infixl "=" 50) +axiomatization equal :: "i \ i \ o" (infixl \=\ 50) where refl [intro]: "x = x" and subst: "x = y \ P x \ P y" @@ -129,11 +129,11 @@ subsection \Quantifiers\ -axiomatization All :: "(i \ o) \ o" (binder "\" 10) +axiomatization All :: "(i \ o) \ o" (binder \\\ 10) where allI [intro]: "(\x. P x) \ \x. P x" and allD [dest]: "\x. P x \ P a" -axiomatization Ex :: "(i \ o) \ o" (binder "\" 10) +axiomatization Ex :: "(i \ o) \ o" (binder \\\ 10) where exI [intro]: "P a \ \x. P x" and exE [elim]: "\x. P x \ (\x. P x \ C) \ C" diff -r 46f59511b7bb -r d97fdabd9e2b src/Pure/Examples/Higher_Order_Logic.thy --- a/src/Pure/Examples/Higher_Order_Logic.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Pure/Examples/Higher_Order_Logic.thy Fri Sep 20 19:51:08 2024 +0200 @@ -27,16 +27,16 @@ instance o :: type .. instance "fun" :: (type, type) type .. -judgment Trueprop :: "o \ prop" ("_" 5) +judgment Trueprop :: "o \ prop" (\_\ 5) section \Minimal logic (axiomatization)\ -axiomatization imp :: "o \ o \ o" (infixr "\" 25) +axiomatization imp :: "o \ o \ o" (infixr \\\ 25) where impI [intro]: "(A \ B) \ A \ B" and impE [dest, trans]: "A \ B \ A \ B" -axiomatization All :: "('a \ o) \ o" (binder "\" 10) +axiomatization All :: "('a \ o) \ o" (binder \\\ 10) where allI [intro]: "(\x. P x) \ \x. P x" and allE [dest]: "\x. P x \ P a" @@ -68,7 +68,7 @@ unfolding True_def .. -definition not :: "o \ o" ("\ _" [40] 40) +definition not :: "o \ o" (\\ _\ [40] 40) where "not \ \A. A \ False" lemma notI [intro]: @@ -91,7 +91,7 @@ lemmas contradiction = notE notE' \ \proof by contradiction in any order\ -definition conj :: "o \ o \ o" (infixr "\" 35) +definition conj :: "o \ o \ o" (infixr \\\ 35) where "A \ B \ \C. (A \ B \ C) \ C" lemma conjI [intro]: @@ -137,7 +137,7 @@ qed -definition disj :: "o \ o \ o" (infixr "\" 30) +definition disj :: "o \ o \ o" (infixr \\\ 30) where "A \ B \ \C. (A \ C) \ (B \ C) \ C" lemma disjI1 [intro]: @@ -190,7 +190,7 @@ qed -definition Ex :: "('a \ o) \ o" (binder "\" 10) +definition Ex :: "('a \ o) \ o" (binder \\\ 10) where "\x. P x \ \C. (\x. P x \ C) \ C" lemma exI [intro]: "P a \ \x. P x" @@ -227,14 +227,14 @@ subsubsection \Extensional equality\ -axiomatization equal :: "'a \ 'a \ o" (infixl "=" 50) +axiomatization equal :: "'a \ 'a \ o" (infixl \=\ 50) where refl [intro]: "x = x" and subst: "x = y \ P x \ P y" -abbreviation not_equal :: "'a \ 'a \ o" (infixl "\" 50) +abbreviation not_equal :: "'a \ 'a \ o" (infixl \\\ 50) where "x \ y \ \ (x = y)" -abbreviation iff :: "o \ o \ o" (infixr "\" 25) +abbreviation iff :: "o \ o \ o" (infixr \\\ 25) where "A \ B \ A = B" axiomatization @@ -401,7 +401,7 @@ axiomatization Eps :: "('a \ o) \ 'a" where someI: "P x \ P (Eps P)" -syntax "_Eps" :: "pttrn \ o \ 'a" ("(3SOME _./ _)" [0, 10] 10) +syntax "_Eps" :: "pttrn \ o \ 'a" (\(3SOME _./ _)\ [0, 10] 10) syntax_consts "_Eps" \ Eps translations "SOME x. P" \ "CONST Eps (\x. P)" diff -r 46f59511b7bb -r d97fdabd9e2b src/Sequents/ILL.thy --- a/src/Sequents/ILL.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Sequents/ILL.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,14 +10,14 @@ consts Trueprop :: "two_seqi" - tens :: "[o, o] \ o" (infixr "><" 35) - limp :: "[o, o] \ o" (infixr "-o" 45) - FShriek :: "o \ o" ("! _" [100] 1000) - lconj :: "[o, o] \ o" (infixr "&&" 35) - ldisj :: "[o, o] \ o" (infixr "++" 35) - zero :: "o" ("0") - top :: "o" ("1") - eye :: "o" ("I") + tens :: "[o, o] \ o" (infixr \><\ 35) + limp :: "[o, o] \ o" (infixr \-o\ 45) + FShriek :: "o \ o" (\! _\ [100] 1000) + lconj :: "[o, o] \ o" (infixr \&&\ 35) + ldisj :: "[o, o] \ o" (infixr \++\ 35) + zero :: "o" (\0\) + top :: "o" (\1\) + eye :: "o" (\I\) (* context manipulation *) @@ -29,9 +29,9 @@ PromAux :: "three_seqi" syntax - "_Trueprop" :: "single_seqe" ("((_)/ \ (_))" [6,6] 5) - "_Context" :: "two_seqe" ("((_)/ :=: (_))" [6,6] 5) - "_PromAux" :: "three_seqe" ("promaux {_||_||_}") + "_Trueprop" :: "single_seqe" (\((_)/ \ (_))\ [6,6] 5) + "_Context" :: "two_seqe" (\((_)/ :=: (_))\ [6,6] 5) + "_PromAux" :: "three_seqe" (\promaux {_||_||_}\) parse_translation \ [(\<^syntax_const>\_Trueprop\, K (single_tr \<^const_syntax>\Trueprop\)), @@ -45,10 +45,10 @@ (\<^const_syntax>\PromAux\, K (three_seq_tr' \<^syntax_const>\_PromAux\))] \ -definition liff :: "[o, o] \ o" (infixr "o-o" 45) +definition liff :: "[o, o] \ o" (infixr \o-o\ 45) where "P o-o Q == (P -o Q) >< (Q -o P)" -definition aneg :: "o\o" ("~_") +definition aneg :: "o\o" (\~_\) where "~A == A -o 0" diff -r 46f59511b7bb -r d97fdabd9e2b src/Sequents/ILL_predlog.thy --- a/src/Sequents/ILL_predlog.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Sequents/ILL_predlog.thy Fri Sep 20 19:51:08 2024 +0200 @@ -5,16 +5,16 @@ typedecl plf consts - conj :: "[plf,plf] \ plf" (infixr "&" 35) - disj :: "[plf,plf] \ plf" (infixr "|" 35) - impl :: "[plf,plf] \ plf" (infixr ">" 35) - eq :: "[plf,plf] \ plf" (infixr "=" 35) - ff :: "plf" ("ff") + conj :: "[plf,plf] \ plf" (infixr \&\ 35) + disj :: "[plf,plf] \ plf" (infixr \|\ 35) + impl :: "[plf,plf] \ plf" (infixr \>\ 35) + eq :: "[plf,plf] \ plf" (infixr \=\ 35) + ff :: "plf" (\ff\) - PL :: "plf \ o" ("[* / _ / *]" [] 100) + PL :: "plf \ o" (\[* / _ / *]\ [] 100) syntax - "_NG" :: "plf \ plf" ("- _ " [50] 55) + "_NG" :: "plf \ plf" (\- _ \ [50] 55) translations "[* A & B *]" \ "[*A*] && [*B*]" diff -r 46f59511b7bb -r d97fdabd9e2b src/Sequents/LK/Nat.thy --- a/src/Sequents/LK/Nat.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Sequents/LK/Nat.thy Fri Sep 20 19:51:08 2024 +0200 @@ -13,7 +13,7 @@ instance nat :: "term" .. axiomatization - Zero :: nat ("0") and + Zero :: nat (\0\) and Suc :: "nat \ nat" and rec :: "[nat, 'a, [nat,'a] \ 'a] \ 'a" where @@ -25,7 +25,7 @@ rec_0: "\ rec(0,a,f) = a" and rec_Suc: "\ rec(Suc(m), a, f) = f(m, rec(m,a,f))" -definition add :: "[nat, nat] \ nat" (infixl "+" 60) +definition add :: "[nat, nat] \ nat" (infixl \+\ 60) where "m + n == rec(m, n, \x y. Suc(y))" diff -r 46f59511b7bb -r d97fdabd9e2b src/Sequents/LK0.thy --- a/src/Sequents/LK0.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Sequents/LK0.thy Fri Sep 20 19:51:08 2024 +0200 @@ -23,24 +23,24 @@ True :: o False :: o - equal :: "['a,'a] \ o" (infixl "=" 50) - Not :: "o \ o" ("\ _" [40] 40) - conj :: "[o,o] \ o" (infixr "\" 35) - disj :: "[o,o] \ o" (infixr "\" 30) - imp :: "[o,o] \ o" (infixr "\" 25) - iff :: "[o,o] \ o" (infixr "\" 25) - The :: "('a \ o) \ 'a" (binder "THE " 10) - All :: "('a \ o) \ o" (binder "\" 10) - Ex :: "('a \ o) \ o" (binder "\" 10) + equal :: "['a,'a] \ o" (infixl \=\ 50) + Not :: "o \ o" (\\ _\ [40] 40) + conj :: "[o,o] \ o" (infixr \\\ 35) + disj :: "[o,o] \ o" (infixr \\\ 30) + imp :: "[o,o] \ o" (infixr \\\ 25) + iff :: "[o,o] \ o" (infixr \\\ 25) + The :: "('a \ o) \ 'a" (binder \THE \ 10) + All :: "('a \ o) \ o" (binder \\\ 10) + Ex :: "('a \ o) \ o" (binder \\\ 10) syntax - "_Trueprop" :: "two_seqe" ("((_)/ \ (_))" [6,6] 5) + "_Trueprop" :: "two_seqe" (\((_)/ \ (_))\ [6,6] 5) parse_translation \[(\<^syntax_const>\_Trueprop\, K (two_seq_tr \<^const_syntax>\Trueprop\))]\ print_translation \[(\<^const_syntax>\Trueprop\, K (two_seq_tr' \<^syntax_const>\_Trueprop\))]\ abbreviation - not_equal (infixl "\" 50) where + not_equal (infixl \\\ 50) where "x \ y \ \ (x = y)" axiomatization where @@ -104,7 +104,7 @@ The: "\$H \ $E, P(a), $F; \x.$H, P(x) \ $E, x=a, $F\ \ $H \ $E, P(THE x. P(x)), $F" -definition If :: "[o, 'a, 'a] \ 'a" ("(if (_)/ then (_)/ else (_))" 10) +definition If :: "[o, 'a, 'a] \ 'a" (\(if (_)/ then (_)/ else (_))\ 10) where "If(P,x,y) \ THE z::'a. (P \ z = x) \ (\ P \ z = y)" diff -r 46f59511b7bb -r d97fdabd9e2b src/Sequents/Modal0.thy --- a/src/Sequents/Modal0.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Sequents/Modal0.thy Fri Sep 20 19:51:08 2024 +0200 @@ -10,14 +10,14 @@ ML_file \modal.ML\ consts - box :: "o\o" ("[]_" [50] 50) - dia :: "o\o" ("<>_" [50] 50) + box :: "o\o" (\[]_\ [50] 50) + dia :: "o\o" (\<>_\ [50] 50) Lstar :: "two_seqi" Rstar :: "two_seqi" syntax - "_Lstar" :: "two_seqe" ("(_)|L>(_)" [6,6] 5) - "_Rstar" :: "two_seqe" ("(_)|R>(_)" [6,6] 5) + "_Lstar" :: "two_seqe" (\(_)|L>(_)\ [6,6] 5) + "_Rstar" :: "two_seqe" (\(_)|R>(_)\ [6,6] 5) ML \ fun star_tr c [s1, s2] = Syntax.const c $ seq_tr s1 $ seq_tr s2; @@ -34,10 +34,10 @@ (\<^const_syntax>\Rstar\, K (star_tr' \<^syntax_const>\_Rstar\))] \ -definition strimp :: "[o,o]\o" (infixr "--<" 25) +definition strimp :: "[o,o]\o" (infixr \--<\ 25) where "P --< Q == [](P \ Q)" -definition streqv :: "[o,o]\o" (infixr ">-<" 25) +definition streqv :: "[o,o]\o" (infixr \>-<\ 25) where "P >-< Q == (P --< Q) \ (Q --< P)" diff -r 46f59511b7bb -r d97fdabd9e2b src/Sequents/S43.thy --- a/src/Sequents/S43.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Sequents/S43.thy Fri Sep 20 19:51:08 2024 +0200 @@ -14,7 +14,7 @@ seq'\seq', seq'\seq', seq'\seq'] \ prop" syntax "_S43pi" :: "[seq, seq, seq, seq, seq, seq] \ prop" - ("S43pi((_);(_);(_);(_);(_);(_))" [] 5) + (\S43pi((_);(_);(_);(_);(_);(_))\ [] 5) parse_translation \ let diff -r 46f59511b7bb -r d97fdabd9e2b src/Sequents/Sequents.thy --- a/src/Sequents/Sequents.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/Sequents/Sequents.thy Fri Sep 20 19:51:08 2024 +0200 @@ -30,14 +30,14 @@ nonterminal seq and seqobj and seqcont syntax - "_SeqEmp" :: seq ("") - "_SeqApp" :: "[seqobj,seqcont] \ seq" ("__") + "_SeqEmp" :: seq (\\) + "_SeqApp" :: "[seqobj,seqcont] \ seq" (\__\) - "_SeqContEmp" :: seqcont ("") - "_SeqContApp" :: "[seqobj,seqcont] \ seqcont" (",/ __") + "_SeqContEmp" :: seqcont (\\) + "_SeqContApp" :: "[seqobj,seqcont] \ seqcont" (\,/ __\) - "_SeqO" :: "o \ seqobj" ("_") - "_SeqId" :: "'a \ seqobj" ("$_") + "_SeqO" :: "o \ seqobj" (\_\) + "_SeqId" :: "'a \ seqobj" (\$_\) type_synonym single_seqe = "[seq,seqobj] \ prop" type_synonym single_seqi = "[seq'\seq',seq'\seq'] \ prop" @@ -52,7 +52,7 @@ syntax (*Constant to allow definitions of SEQUENCES of formulas*) - "_Side" :: "seq\(seq'\seq')" ("<<(_)>>") + "_Side" :: "seq\(seq'\seq')" (\<<(_)>>\) ML \