standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
--- 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]\<Rightarrow>prop" (infixl "\<longlongrightarrow>" 20)
+ Eval :: "[i,i]\<Rightarrow>prop" (infixl \<open>\<longlongrightarrow>\<close> 20)
(*** Bisimulations for pre-order and equality ***)
- po :: "['a,'a]\<Rightarrow>o" (infixl "[=" 50)
+ po :: "['a,'a]\<Rightarrow>o" (infixl \<open>[=\<close> 50)
(*** Term Formers ***)
true :: "i"
false :: "i"
- pair :: "[i,i]\<Rightarrow>i" ("(1<_,/_>)")
- lambda :: "(i\<Rightarrow>i)\<Rightarrow>i" (binder "lam " 55)
+ pair :: "[i,i]\<Rightarrow>i" (\<open>(1<_,/_>)\<close>)
+ lambda :: "(i\<Rightarrow>i)\<Rightarrow>i" (binder \<open>lam \<close> 55)
"case" :: "[i,i,i,[i,i]\<Rightarrow>i,(i\<Rightarrow>i)\<Rightarrow>i]\<Rightarrow>i"
- "apply" :: "[i,i]\<Rightarrow>i" (infixl "`" 56)
+ "apply" :: "[i,i]\<Rightarrow>i" (infixl \<open>`\<close> 56)
bot :: "i"
(******* EVALUATION SEMANTICS *******)
--- 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 \<open>Set comprehension and membership\<close>
axiomatization Collect :: "['a \<Rightarrow> o] \<Rightarrow> 'a set"
- and mem :: "['a, 'a set] \<Rightarrow> o" (infixl ":" 50)
+ and mem :: "['a, 'a set] \<Rightarrow> o" (infixl \<open>:\<close> 50)
where mem_Collect_iff: "(a : Collect(P)) \<longleftrightarrow> P(a)"
and set_extension: "A = B \<longleftrightarrow> (ALL x. x:A \<longleftrightarrow> x:B)"
syntax
- "_Coll" :: "[idt, o] \<Rightarrow> 'a set" ("(1{_./ _})")
+ "_Coll" :: "[idt, o] \<Rightarrow> 'a set" (\<open>(1{_./ _})\<close>)
syntax_consts
"_Coll" == Collect
translations
@@ -50,8 +50,8 @@
where "Bex(A, P) == EX x. x:A \<and> P(x)"
syntax
- "_Ball" :: "[idt, 'a set, o] \<Rightarrow> o" ("(ALL _:_./ _)" [0, 0, 0] 10)
- "_Bex" :: "[idt, 'a set, o] \<Rightarrow> o" ("(EX _:_./ _)" [0, 0, 0] 10)
+ "_Ball" :: "[idt, 'a set, o] \<Rightarrow> o" (\<open>(ALL _:_./ _)\<close> [0, 0, 0] 10)
+ "_Bex" :: "[idt, 'a set, o] \<Rightarrow> o" (\<open>(EX _:_./ _)\<close> [0, 0, 0] 10)
syntax_consts
"_Ball" == Ball and
"_Bex" == Bex
@@ -96,22 +96,22 @@
subsection \<open>Further operations\<close>
-definition subset :: "['a set, 'a set] \<Rightarrow> o" (infixl "<=" 50)
+definition subset :: "['a set, 'a set] \<Rightarrow> o" (infixl \<open><=\<close> 50)
where "A <= B == ALL x:A. x:B"
definition mono :: "['a set \<Rightarrow> 'b set] \<Rightarrow> o"
where "mono(f) == (ALL A B. A <= B \<longrightarrow> f(A) <= f(B))"
-definition singleton :: "'a \<Rightarrow> 'a set" ("{_}")
+definition singleton :: "'a \<Rightarrow> 'a set" (\<open>{_}\<close>)
where "{a} == {x. x=a}"
-definition empty :: "'a set" ("{}")
+definition empty :: "'a set" (\<open>{}\<close>)
where "{} == {x. False}"
-definition Un :: "['a set, 'a set] \<Rightarrow> 'a set" (infixl "Un" 65)
+definition Un :: "['a set, 'a set] \<Rightarrow> 'a set" (infixl \<open>Un\<close> 65)
where "A Un B == {x. x:A | x:B}"
-definition Int :: "['a set, 'a set] \<Rightarrow> 'a set" (infixl "Int" 70)
+definition Int :: "['a set, 'a set] \<Rightarrow> 'a set" (infixl \<open>Int\<close> 70)
where "A Int B == {x. x:A \<and> x:B}"
definition Compl :: "('a set) \<Rightarrow> 'a set"
@@ -127,8 +127,8 @@
where "UNION(A, B) == {y. EX x:A. y: B(x)}"
syntax
- "_INTER" :: "[idt, 'a set, 'b set] \<Rightarrow> 'b set" ("(INT _:_./ _)" [0, 0, 0] 10)
- "_UNION" :: "[idt, 'a set, 'b set] \<Rightarrow> 'b set" ("(UN _:_./ _)" [0, 0, 0] 10)
+ "_INTER" :: "[idt, 'a set, 'b set] \<Rightarrow> 'b set" (\<open>(INT _:_./ _)\<close> [0, 0, 0] 10)
+ "_UNION" :: "[idt, 'a set, 'b set] \<Rightarrow> 'b set" (\<open>(UN _:_./ _)\<close> [0, 0, 0] 10)
syntax_consts
"_INTER" == INTER and
"_UNION" == UNION
--- 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]\<Rightarrow>i" ("(3if _/ then _/ else _)" [0,0,60] 60)
+definition "if" :: "[i,i,i]\<Rightarrow>i" (\<open>(3if _/ then _/ else _)\<close> [0,0,60] 60)
where "if b then t else u == case(b, t, u, \<lambda> x y. bot, \<lambda>v. bot)"
definition inl :: "i\<Rightarrow>i"
@@ -48,7 +48,7 @@
definition "let1" :: "[i,i\<Rightarrow>i]\<Rightarrow>i"
where let_def: "let1(t, f) == case(t,f(true),f(false), \<lambda>x y. f(<x,y>), \<lambda>u. f(lam x. u(x)))"
-syntax "_let1" :: "[idt,i,i]\<Rightarrow>i" ("(3let _ be _/ in _)" [0,0,60] 60)
+syntax "_let1" :: "[idt,i,i]\<Rightarrow>i" (\<open>(3let _ be _/ in _)\<close> [0,0,60] 60)
syntax_consts "_let1" == let1
translations "let x be a in e" == "CONST let1(a, \<lambda>x. e)"
@@ -65,9 +65,9 @@
\<lambda>g'. f(\<lambda>x y z. g'(<x,<y,z>>)))"
syntax
- "_letrec" :: "[idt,idt,i,i]\<Rightarrow>i" ("(3letrec _ _ be _/ in _)" [0,0,0,60] 60)
- "_letrec2" :: "[idt,idt,idt,i,i]\<Rightarrow>i" ("(3letrec _ _ _ be _/ in _)" [0,0,0,0,60] 60)
- "_letrec3" :: "[idt,idt,idt,idt,i,i]\<Rightarrow>i" ("(3letrec _ _ _ _ be _/ in _)" [0,0,0,0,0,60] 60)
+ "_letrec" :: "[idt,idt,i,i]\<Rightarrow>i" (\<open>(3letrec _ _ be _/ in _)\<close> [0,0,0,60] 60)
+ "_letrec2" :: "[idt,idt,idt,i,i]\<Rightarrow>i" (\<open>(3letrec _ _ _ be _/ in _)\<close> [0,0,0,0,60] 60)
+ "_letrec3" :: "[idt,idt,idt,idt,i,i]\<Rightarrow>i" (\<open>(3letrec _ _ _ _ be _/ in _)\<close> [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]\<Rightarrow>i]\<Rightarrow>i"
where "nrec(n,b,c) == letrec g x be ncase(x, b, \<lambda>y. c(y,g(y))) in g(n)"
-definition nil :: "i" ("([])")
+definition nil :: "i" (\<open>([])\<close>)
where "[] == inl(one)"
-definition cons :: "[i,i]\<Rightarrow>i" (infixr "$" 80)
+definition cons :: "[i,i]\<Rightarrow>i" (infixr \<open>$\<close> 80)
where "h$t == inr(<h,t>)"
definition lcase :: "[i,i,[i,i]\<Rightarrow>i]\<Rightarrow>i"
@@ -143,7 +143,7 @@
definition lrec :: "[i,i,[i,i,i]\<Rightarrow>i]\<Rightarrow>i"
where "lrec(l,b,c) == letrec g x be lcase(x, b, \<lambda>h t. c(h,t,g(t))) in g(l)"
-definition napply :: "[i\<Rightarrow>i,i,i]\<Rightarrow>i" ("(_ ^ _ ` _)" [56,56,56] 56)
+definition napply :: "[i\<Rightarrow>i,i,i]\<Rightarrow>i" (\<open>(_ ^ _ ` _)\<close> [56,56,56] 56)
where "f ^n` a == nrec(n,a,\<lambda>x g. f(g))"
lemmas simp_can_defs = one_def inl_def inr_def
--- 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 = <x,x>}"
-definition relcomp :: "[i set,i set] \<Rightarrow> i set" (infixr "O" 60) (*composition of relations*)
+definition relcomp :: "[i set,i set] \<Rightarrow> i set" (infixr \<open>O\<close> 60) (*composition of relations*)
where "r O s == {xz. EX x y z. xz = <x,z> \<and> <x,y>:s \<and> <y,z>:r}"
-definition rtrancl :: "i set \<Rightarrow> i set" ("(_^*)" [100] 100)
+definition rtrancl :: "i set \<Rightarrow> i set" (\<open>(_^*)\<close> [100] 100)
where "r^* == lfp(\<lambda>s. id Un (r O s))"
-definition trancl :: "i set \<Rightarrow> i set" ("(_^+)" [100] 100)
+definition trancl :: "i set \<Rightarrow> i set" (\<open>(_^+)\<close> [100] 100)
where "r^+ == r O rtrancl(r)"
--- 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 \<and> P(x)}"
syntax
- "_Subtype" :: "[idt, 'a set, o] \<Rightarrow> 'a set" ("(1{_: _ ./ _})")
+ "_Subtype" :: "[idt, 'a set, o] \<Rightarrow> 'a set" (\<open>(1{_: _ ./ _})\<close>)
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] \<Rightarrow> i set" (infixr "+" 55)
+definition Plus :: "[i set, i set] \<Rightarrow> i set" (infixr \<open>+\<close> 55)
where "A+B == {x. (EX a:A. x=inl(a)) | (EX b:B. x=inr(b))}"
definition Pi :: "[i set, i \<Rightarrow> i set] \<Rightarrow> i set"
@@ -35,10 +35,10 @@
where "Sigma(A,B) == {x. EX a:A. EX b:B(a).x=<a,b>}"
syntax
- "_Pi" :: "[idt, i set, i set] \<Rightarrow> i set" ("(3PROD _:_./ _)" [0,0,60] 60)
- "_Sigma" :: "[idt, i set, i set] \<Rightarrow> i set" ("(3SUM _:_./ _)" [0,0,60] 60)
- "_arrow" :: "[i set, i set] \<Rightarrow> i set" ("(_ ->/ _)" [54, 53] 53)
- "_star" :: "[i set, i set] \<Rightarrow> i set" ("(_ */ _)" [56, 55] 55)
+ "_Pi" :: "[idt, i set, i set] \<Rightarrow> i set" (\<open>(3PROD _:_./ _)\<close> [0,0,60] 60)
+ "_Sigma" :: "[idt, i set, i set] \<Rightarrow> i set" (\<open>(3SUM _:_./ _)\<close> [0,0,60] 60)
+ "_arrow" :: "[i set, i set] \<Rightarrow> i set" (\<open>(_ ->/ _)\<close> [54, 53] 53)
+ "_star" :: "[i set, i set] \<Rightarrow> i set" (\<open>(_ */ _)\<close> [56, 55] 55)
syntax_consts
"_Pi" "_arrow" \<rightleftharpoons> Pi and
"_Sigma" "_star" \<rightleftharpoons> Sigma
@@ -67,13 +67,13 @@
where "ILists(A) == gfp(\<lambda>X.{} + A*X)"
-definition TAll :: "(i set \<Rightarrow> i set) \<Rightarrow> i set" (binder "TALL " 55)
+definition TAll :: "(i set \<Rightarrow> i set) \<Rightarrow> i set" (binder \<open>TALL \<close> 55)
where "TALL X. B(X) == Inter({X. EX Y. X=B(Y)})"
-definition TEx :: "(i set \<Rightarrow> i set) \<Rightarrow> i set" (binder "TEX " 55)
+definition TEx :: "(i set \<Rightarrow> i set) \<Rightarrow> i set" (binder \<open>TEX \<close> 55)
where "TEX X. B(X) == Union({X. EX Y. X=B(Y)})"
-definition Lift :: "i set \<Rightarrow> i set" ("(3[_])")
+definition Lift :: "i set \<Rightarrow> i set" (\<open>(3[_])\<close>)
where "[A] == A Un {bot}"
definition SPLIT :: "[i, [i, i] \<Rightarrow> i set] \<Rightarrow> i set"
--- 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\<Rightarrow>i,i set] \<Rightarrow> i set"
where "wmap(f,R) == {p. EX x y. p=<x,y> \<and> <f(x),f(y)> : R}"
-definition lex :: "[i set,i set] => i set" (infixl "**" 70)
+definition lex :: "[i set,i set] => i set" (infixl \<open>**\<close> 70)
where "ra**rb == {p. EX a a' b b'. p = <<a,b>,<a',b'>> \<and> (<a,a'> : ra | (a=a' \<and> <b,b'> : rb))}"
definition NatPR :: "i set"
--- 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\<Rightarrow>i,i]\<Rightarrow>i"
where "map(f,l) == lrec(l, [], \<lambda>x xs g. f(x)$g)"
-definition comp :: "[i\<Rightarrow>i,i\<Rightarrow>i]\<Rightarrow>i\<Rightarrow>i" (infixr "\<circ>" 55)
+definition comp :: "[i\<Rightarrow>i,i\<Rightarrow>i]\<Rightarrow>i\<Rightarrow>i" (infixr \<open>\<circ>\<close> 55)
where "f \<circ> g == (\<lambda>x. f(g(x)))"
-definition append :: "[i,i]\<Rightarrow>i" (infixr "@" 55)
+definition append :: "[i,i]\<Rightarrow>i" (infixr \<open>@\<close> 55)
where "l @ m == lrec(l, m, \<lambda>x xs g. x$g)"
-axiomatization member :: "[i,i]\<Rightarrow>i" (infixr "mem" 55) (* FIXME dangling eq *)
+axiomatization member :: "[i,i]\<Rightarrow>i" (infixr \<open>mem\<close> 55) (* FIXME dangling eq *)
where member_ax: "a mem l == lrec(l, false, \<lambda>h t g. if eq(a,h) then true else g)"
definition filter :: "[i,i]\<Rightarrow>i"
--- 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\<Rightarrow>i"
where "not(b) == if b then false else true"
-definition add :: "[i,i]\<Rightarrow>i" (infixr "#+" 60)
+definition add :: "[i,i]\<Rightarrow>i" (infixr \<open>#+\<close> 60)
where "a #+ b == nrec(a, b, \<lambda>x g. succ(g))"
-definition mult :: "[i,i]\<Rightarrow>i" (infixr "#*" 60)
+definition mult :: "[i,i]\<Rightarrow>i" (infixr \<open>#*\<close> 60)
where "a #* b == nrec(a, zero, \<lambda>x g. b #+ g)"
-definition sub :: "[i,i]\<Rightarrow>i" (infixr "#-" 60)
+definition sub :: "[i,i]\<Rightarrow>i" (infixr \<open>#-\<close> 60)
where
"a #- b ==
letrec sub x y be ncase(y, x, \<lambda>yy. ncase(x, zero, \<lambda>xx. sub(xx,yy)))
in sub(a,b)"
-definition le :: "[i,i]\<Rightarrow>i" (infixr "#<=" 60)
+definition le :: "[i,i]\<Rightarrow>i" (infixr \<open>#<=\<close> 60)
where
"a #<= b ==
letrec le x y be ncase(x, true, \<lambda>xx. ncase(y, false, \<lambda>yy. le(xx,yy)))
in le(a,b)"
-definition lt :: "[i,i]\<Rightarrow>i" (infixr "#<" 60)
+definition lt :: "[i,i]\<Rightarrow>i" (infixr \<open>#<\<close> 60)
where "a #< b == not(b #<= a)"
-definition div :: "[i,i]\<Rightarrow>i" (infixr "##" 60)
+definition div :: "[i,i]\<Rightarrow>i" (infixr \<open>##\<close> 60)
where
"a ## b ==
letrec div x y be if x #< y then zero else succ(div(x#-y,y))
--- 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
\<comment> \<open>Judgments\<close>
- Type :: "t \<Rightarrow> prop" ("(_ type)" [10] 5)
- Eqtype :: "[t,t]\<Rightarrow>prop" ("(_ =/ _)" [10,10] 5)
- Elem :: "[i, t]\<Rightarrow>prop" ("(_ /: _)" [10,10] 5)
- Eqelem :: "[i,i,t]\<Rightarrow>prop" ("(_ =/ _ :/ _)" [10,10,10] 5)
- Reduce :: "[i,i]\<Rightarrow>prop" ("Reduce[_,_]")
+ Type :: "t \<Rightarrow> prop" (\<open>(_ type)\<close> [10] 5)
+ Eqtype :: "[t,t]\<Rightarrow>prop" (\<open>(_ =/ _)\<close> [10,10] 5)
+ Elem :: "[i, t]\<Rightarrow>prop" (\<open>(_ /: _)\<close> [10,10] 5)
+ Eqelem :: "[i,i,t]\<Rightarrow>prop" (\<open>(_ =/ _ :/ _)\<close> [10,10,10] 5)
+ Reduce :: "[i,i]\<Rightarrow>prop" (\<open>Reduce[_,_]\<close>)
\<comment> \<open>Types for truth values\<close>
F :: "t"
T :: "t" \<comment> \<open>\<open>F\<close> is empty, \<open>T\<close> contains one element\<close>
@@ -30,24 +30,24 @@
tt :: "i"
\<comment> \<open>Natural numbers\<close>
N :: "t"
- Zero :: "i" ("0")
+ Zero :: "i" (\<open>0\<close>)
succ :: "i\<Rightarrow>i"
rec :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow> i"
\<comment> \<open>Binary sum\<close>
- Plus :: "[t,t]\<Rightarrow>t" (infixr "+" 40)
+ Plus :: "[t,t]\<Rightarrow>t" (infixr \<open>+\<close> 40)
inl :: "i\<Rightarrow>i"
inr :: "i\<Rightarrow>i"
"when" :: "[i, i\<Rightarrow>i, i\<Rightarrow>i]\<Rightarrow>i"
\<comment> \<open>General sum and binary product\<close>
Sum :: "[t, i\<Rightarrow>t]\<Rightarrow>t"
- pair :: "[i,i]\<Rightarrow>i" ("(1<_,/_>)")
+ pair :: "[i,i]\<Rightarrow>i" (\<open>(1<_,/_>)\<close>)
fst :: "i\<Rightarrow>i"
snd :: "i\<Rightarrow>i"
split :: "[i, [i,i]\<Rightarrow>i] \<Rightarrow>i"
\<comment> \<open>General product and function space\<close>
Prod :: "[t, i\<Rightarrow>t]\<Rightarrow>t"
- lambda :: "(i \<Rightarrow> i) \<Rightarrow> i" (binder "\<^bold>\<lambda>" 10)
- app :: "[i,i]\<Rightarrow>i" (infixl "`" 60)
+ lambda :: "(i \<Rightarrow> i) \<Rightarrow> i" (binder \<open>\<^bold>\<lambda>\<close> 10)
+ app :: "[i,i]\<Rightarrow>i" (infixl \<open>`\<close> 60)
\<comment> \<open>Equality type\<close>
Eq :: "[t,i,i]\<Rightarrow>t"
eq :: "i"
@@ -56,8 +56,8 @@
must be introduced after the judgment forms.\<close>
syntax
- "_PROD" :: "[idt,t,t]\<Rightarrow>t" ("(3\<Prod>_:_./ _)" 10)
- "_SUM" :: "[idt,t,t]\<Rightarrow>t" ("(3\<Sum>_:_./ _)" 10)
+ "_PROD" :: "[idt,t,t]\<Rightarrow>t" (\<open>(3\<Prod>_:_./ _)\<close> 10)
+ "_SUM" :: "[idt,t,t]\<Rightarrow>t" (\<open>(3\<Sum>_:_./ _)\<close> 10)
syntax_consts
"_PROD" \<rightleftharpoons> Prod and
"_SUM" \<rightleftharpoons> Sum
@@ -65,10 +65,10 @@
"\<Prod>x:A. B" \<rightleftharpoons> "CONST Prod(A, \<lambda>x. B)"
"\<Sum>x:A. B" \<rightleftharpoons> "CONST Sum(A, \<lambda>x. B)"
-abbreviation Arrow :: "[t,t]\<Rightarrow>t" (infixr "\<longrightarrow>" 30)
+abbreviation Arrow :: "[t,t]\<Rightarrow>t" (infixr \<open>\<longrightarrow>\<close> 30)
where "A \<longrightarrow> B \<equiv> \<Prod>_:A. B"
-abbreviation Times :: "[t,t]\<Rightarrow>t" (infixr "\<times>" 50)
+abbreviation Times :: "[t,t]\<Rightarrow>t" (infixr \<open>\<times>\<close> 50)
where "A \<times> B \<equiv> \<Sum>_:A. B"
text \<open>
@@ -547,22 +547,22 @@
subsection \<open>Arithmetic operators and their definitions\<close>
-definition add :: "[i,i]\<Rightarrow>i" (infixr "#+" 65)
+definition add :: "[i,i]\<Rightarrow>i" (infixr \<open>#+\<close> 65)
where "a#+b \<equiv> rec(a, b, \<lambda>u v. succ(v))"
-definition diff :: "[i,i]\<Rightarrow>i" (infixr "-" 65)
+definition diff :: "[i,i]\<Rightarrow>i" (infixr \<open>-\<close> 65)
where "a-b \<equiv> rec(b, a, \<lambda>u v. rec(v, 0, \<lambda>x y. x))"
-definition absdiff :: "[i,i]\<Rightarrow>i" (infixr "|-|" 65)
+definition absdiff :: "[i,i]\<Rightarrow>i" (infixr \<open>|-|\<close> 65)
where "a|-|b \<equiv> (a-b) #+ (b-a)"
-definition mult :: "[i,i]\<Rightarrow>i" (infixr "#*" 70)
+definition mult :: "[i,i]\<Rightarrow>i" (infixr \<open>#*\<close> 70)
where "a#*b \<equiv> rec(a, 0, \<lambda>u v. b #+ v)"
-definition mod :: "[i,i]\<Rightarrow>i" (infixr "mod" 70)
+definition mod :: "[i,i]\<Rightarrow>i" (infixr \<open>mod\<close> 70)
where "a mod b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(v) |-| b, 0, \<lambda>x y. succ(v)))"
-definition div :: "[i,i]\<Rightarrow>i" (infixr "div" 70)
+definition div :: "[i,i]\<Rightarrow>i" (infixr \<open>div\<close> 70)
where "a div b \<equiv> rec(a, 0, \<lambda>u v. rec(succ(u) mod b, succ(v), \<lambda>x y. v))"
lemmas arith_defs = add_def diff_def absdiff_def mult_def mod_def div_def
--- 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] \<Rightarrow> prop" and
MT_context :: "context" and
Context :: "[typing, context] \<Rightarrow> context" and
- star :: "term" ("*") and
- box :: "term" ("\<box>") and
- app :: "[term, term] \<Rightarrow> term" (infixl "\<cdot>" 20) and
+ star :: "term" (\<open>*\<close>) and
+ box :: "term" (\<open>\<box>\<close>) and
+ app :: "[term, term] \<Rightarrow> term" (infixl \<open>\<cdot>\<close> 20) and
Has_type :: "[term, term] \<Rightarrow> typing"
nonterminal context' and typing'
syntax
- "_Trueprop" :: "[context', typing'] \<Rightarrow> prop" ("(_/ \<turnstile> _)")
- "_Trueprop1" :: "typing' \<Rightarrow> prop" ("(_)")
- "" :: "id \<Rightarrow> context'" ("_")
- "" :: "var \<Rightarrow> context'" ("_")
- "_MT_context" :: "context'" ("")
- "_Context" :: "[typing', context'] \<Rightarrow> context'" ("_ _")
- "_Has_type" :: "[term, term] \<Rightarrow> typing'" ("(_:/ _)" [0, 0] 5)
- "_Lam" :: "[idt, term, term] \<Rightarrow> term" ("(3\<^bold>\<lambda>_:_./ _)" [0, 0, 0] 10)
- "_Pi" :: "[idt, term, term] \<Rightarrow> term" ("(3\<Prod>_:_./ _)" [0, 0] 10)
- "_arrow" :: "[term, term] \<Rightarrow> term" (infixr "\<rightarrow>" 10)
+ "_Trueprop" :: "[context', typing'] \<Rightarrow> prop" (\<open>(_/ \<turnstile> _)\<close>)
+ "_Trueprop1" :: "typing' \<Rightarrow> prop" (\<open>(_)\<close>)
+ "" :: "id \<Rightarrow> context'" (\<open>_\<close>)
+ "" :: "var \<Rightarrow> context'" (\<open>_\<close>)
+ "_MT_context" :: "context'" (\<open>\<close>)
+ "_Context" :: "[typing', context'] \<Rightarrow> context'" (\<open>_ _\<close>)
+ "_Has_type" :: "[term, term] \<Rightarrow> typing'" (\<open>(_:/ _)\<close> [0, 0] 5)
+ "_Lam" :: "[idt, term, term] \<Rightarrow> term" (\<open>(3\<^bold>\<lambda>_:_./ _)\<close> [0, 0, 0] 10)
+ "_Pi" :: "[idt, term, term] \<Rightarrow> term" (\<open>(3\<Prod>_:_./ _)\<close> [0, 0] 10)
+ "_arrow" :: "[term, term] \<Rightarrow> term" (infixr \<open>\<rightarrow>\<close> 10)
syntax_consts
"_Trueprop" \<rightleftharpoons> Trueprop and
"_MT_context" \<rightleftharpoons> MT_context and
--- 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>")
- "_alpha_ofsort" :: "sort \<Rightarrow> type" ("\<alpha>' ::_" [0] 1000)
- "_beta" :: "type" ("\<beta>")
- "_beta_ofsort" :: "sort \<Rightarrow> type" ("\<beta>' ::_" [0] 1000)
+ "_alpha" :: "type" (\<open>\<alpha>\<close>)
+ "_alpha_ofsort" :: "sort \<Rightarrow> type" (\<open>\<alpha>' ::_\<close> [0] 1000)
+ "_beta" :: "type" (\<open>\<beta>\<close>)
+ "_beta_ofsort" :: "sort \<Rightarrow> type" (\<open>\<beta>' ::_\<close> [0] 1000)
parse_ast_translation \<open>
let
--- 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 @@
\<close>
class %quote semigroup =
- fixes mult :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<otimes>" 70)
+ fixes mult :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<otimes>\<close> 70)
assumes assoc: "(x \<otimes> y) \<otimes> z = x \<otimes> (y \<otimes> z)"
class %quote monoid = semigroup +
- fixes neutral :: 'a ("\<one>")
+ fixes neutral :: 'a (\<open>\<one>\<close>)
assumes neutl: "\<one> \<otimes> x = x"
and neutr: "x \<otimes> \<one> = x"
--- 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 \<open>../more_antiquote.ML\<close>
no_syntax (output)
- "_constrain" :: "logic => type => logic" ("_::_" [4, 0] 3)
- "_constrain" :: "prop' => type => prop'" ("_::_" [4, 0] 3)
+ "_constrain" :: "logic => type => logic" (\<open>_::_\<close> [4, 0] 3)
+ "_constrain" :: "prop' => type => prop'" (\<open>_::_\<close> [4, 0] 3)
syntax (output)
- "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3)
- "_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3)
+ "_constrain" :: "logic => type => logic" (\<open>_ :: _\<close> [4, 0] 3)
+ "_constrain" :: "prop' => type => prop'" (\<open>_ :: _\<close> [4, 0] 3)
declare [[default_code_width = 74]]
--- 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 (\<open>[]\<close>) and
+ Cons (infixr \<open>#\<close> 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 \<open>*\<close> 20) = Pair 'a 'b
text \<open>\blankline\<close>
datatype (set: 'a) list =
- null: Nil ("[]")
- | Cons (hd: 'a) (tl: "'a list") (infixr "#" 65)
+ null: Nil (\<open>[]\<close>)
+ | Cons (hd: 'a) (tl: "'a list") (infixr \<open>#\<close> 65)
for
map: map
rel: list_all2
@@ -453,7 +453,7 @@
Incidentally, this is how the traditional syntax can be set up:
\<close>
- syntax "_list" :: "list_args \<Rightarrow> 'a list" ("[(_)]")
+ syntax "_list" :: "list_args \<Rightarrow> 'a list" (\<open>[(_)]\<close>)
text \<open>\blankline\<close>
--- 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 \<Rightarrow> prop" ("_" 5)
+judgment Trueprop :: "o \<Rightarrow> prop" (\<open>_\<close> 5)
text \<open>
Note that the object-logic judgment is implicit in the syntax: writing
@@ -35,7 +35,7 @@
\<open>B x\<close> against expressions containing occurrences of \<open>x\<close>.
\<close>
-axiomatization equal :: "i \<Rightarrow> i \<Rightarrow> o" (infix "=" 50)
+axiomatization equal :: "i \<Rightarrow> i \<Rightarrow> o" (infix \<open>=\<close> 50)
where refl [intro]: "x = x"
and subst [elim]: "x = y \<Longrightarrow> B x \<Longrightarrow> B y"
@@ -87,9 +87,9 @@
\<close>
locale group =
- fixes prod :: "i \<Rightarrow> i \<Rightarrow> i" (infix "\<circ>" 70)
- and inv :: "i \<Rightarrow> i" ("(_\<inverse>)" [1000] 999)
- and unit :: i ("1")
+ fixes prod :: "i \<Rightarrow> i \<Rightarrow> i" (infix \<open>\<circ>\<close> 70)
+ and inv :: "i \<Rightarrow> i" (\<open>(_\<inverse>)\<close> [1000] 999)
+ and unit :: i (\<open>1\<close>)
assumes assoc: "(x \<circ> y) \<circ> z = x \<circ> (y \<circ> z)"
and left_unit: "1 \<circ> x = x"
and left_inv: "x\<inverse> \<circ> x = 1"
@@ -167,16 +167,16 @@
Gentzen's system of Natural Deduction \<^cite>\<open>"Gentzen:1935"\<close>.
\<close>
-axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longrightarrow>" 25)
+axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<longrightarrow>\<close> 25)
where impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B"
and impD [dest]: "(A \<longrightarrow> B) \<Longrightarrow> A \<Longrightarrow> B"
-axiomatization disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30)
+axiomatization disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<or>\<close> 30)
where disjI\<^sub>1 [intro]: "A \<Longrightarrow> A \<or> B"
and disjI\<^sub>2 [intro]: "B \<Longrightarrow> A \<or> B"
and disjE [elim]: "A \<or> B \<Longrightarrow> (A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C"
-axiomatization conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35)
+axiomatization conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<and>\<close> 35)
where conjI [intro]: "A \<Longrightarrow> B \<Longrightarrow> A \<and> B"
and conjD\<^sub>1: "A \<and> B \<Longrightarrow> A"
and conjD\<^sub>2: "A \<and> B \<Longrightarrow> B"
@@ -223,10 +223,10 @@
that is trivially true.
\<close>
-axiomatization false :: o ("\<bottom>")
+axiomatization false :: o (\<open>\<bottom>\<close>)
where falseE [elim]: "\<bottom> \<Longrightarrow> A"
-definition true :: o ("\<top>")
+definition true :: o (\<open>\<top>\<close>)
where "\<top> \<equiv> \<bottom> \<longrightarrow> \<bottom>"
theorem trueI [intro]: \<top>
@@ -237,7 +237,7 @@
Now negation represents an implication towards absurdity:
\<close>
-definition not :: "o \<Rightarrow> o" ("\<not> _" [40] 40)
+definition not :: "o \<Rightarrow> o" (\<open>\<not> _\<close> [40] 40)
where "\<not> A \<equiv> A \<longrightarrow> \<bottom>"
theorem notI [intro]:
@@ -320,11 +320,11 @@
notation turns \<open>All (\<lambda>x. B x)\<close> into \<open>\<forall>x. B x\<close> etc.
\<close>
-axiomatization All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10)
+axiomatization All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<forall>\<close> 10)
where allI [intro]: "(\<And>x. B x) \<Longrightarrow> \<forall>x. B x"
and allD [dest]: "(\<forall>x. B x) \<Longrightarrow> B a"
-axiomatization Ex :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10)
+axiomatization Ex :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>\<close> 10)
where exI [intro]: "B a \<Longrightarrow> (\<exists>x. B x)"
and exE [elim]: "(\<exists>x. B x) \<Longrightarrow> (\<And>x. B x \<Longrightarrow> C) \<Longrightarrow> C"
--- 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 @@
\<close>
experiment
- fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<bullet>" 60)
+ fixes f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix \<open>\<bullet>\<close> 60)
assumes assoc: "(x \<bullet> y) \<bullet> z = x \<bullet> (y \<bullet> z)"
assumes commute: "x \<bullet> y = y \<bullet> x"
begin
--- 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 \<open>The accessible part of a relation is defined as follows:\<close>
inductive acc :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool"
- for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<prec>" 50)
+ for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open>\<prec>\<close> 50)
where acc: "(\<And>y. y \<prec> x \<Longrightarrow> acc r y) \<Longrightarrow> acc r x"
(*<*)end(*>*)
--- 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 \<open>Notation:\<close>
- write x ("***")
+ write x (\<open>***\<close>)
end
@@ -562,7 +562,7 @@
notepad
begin
- write Trueprop ("Tr")
+ write Trueprop (\<open>Tr\<close>)
thm conjI
thm impI
--- 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 @@
\<close>
locale partial_order =
- fixes le :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sqsubseteq>" 50)
+ fixes le :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>\<sqsubseteq>\<close> 50)
assumes refl [intro, simp]: "x \<sqsubseteq> x"
and anti_sym [intro]: "\<lbrakk> x \<sqsubseteq> y; y \<sqsubseteq> x \<rbrakk> \<Longrightarrow> x = y"
and trans [trans]: "\<lbrakk> x \<sqsubseteq> y; y \<sqsubseteq> z \<rbrakk> \<Longrightarrow> x \<sqsubseteq> z"
@@ -141,7 +141,7 @@
\<close>
definition (in partial_order)
- less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sqsubset>" 50)
+ less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>\<sqsubset>\<close> 50)
where "(x \<sqsubset> y) = (x \<sqsubseteq> y \<and> x \<noteq> y)"
text (in partial_order) \<open>The strict order \<open>less\<close> with infix
@@ -321,9 +321,9 @@
together with an example theorem.\<close>
definition
- meet (infixl "\<sqinter>" 70) where "x \<sqinter> y = (THE inf. is_inf x y inf)"
+ meet (infixl \<open>\<sqinter>\<close> 70) where "x \<sqinter> y = (THE inf. is_inf x y inf)"
definition
- join (infixl "\<squnion>" 65) where "x \<squnion> y = (THE sup. is_sup x y sup)"
+ join (infixl \<open>\<squnion>\<close> 65) where "x \<squnion> y = (THE sup. is_sup x y sup)"
lemma %invisible meet_equality [elim?]: "is_inf x y i \<Longrightarrow> x \<sqinter> y = i"
proof (unfold meet_def)
--- 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 "\<sqsubseteq>" 50) and le' (infixl "\<preceq>" 50) +
+ for le (infixl \<open>\<sqsubseteq>\<close> 50) and le' (infixl \<open>\<preceq>\<close> 50) +
fixes \<phi>
assumes hom_le: "x \<sqsubseteq> y \<Longrightarrow> \<phi> x \<preceq> \<phi> y"
@@ -222,7 +222,7 @@
available for the original instance it was declared for. We may
introduce infix syntax for \<open>le'.less\<close> with the following declaration:\<close>
- notation (in order_preserving) le'.less (infixl "\<prec>" 50)
+ notation (in order_preserving) le'.less (infixl \<open>\<prec>\<close> 50)
text (in order_preserving) \<open>Now the theorem is displayed nicely as
@{thm [source] le'.less_le_trans}:
@@ -279,7 +279,7 @@
join.\<close>
locale lattice_hom =
- le: lattice + le': lattice le' for le' (infixl "\<preceq>" 50) +
+ le: lattice + le': lattice le' for le' (infixl \<open>\<preceq>\<close> 50) +
fixes \<phi>
assumes hom_meet: "\<phi> (x \<sqinter> y) = le'.meet (\<phi> x) (\<phi> y)"
and hom_join: "\<phi> (x \<squnion> y) = le'.join (\<phi> x) (\<phi> y)"
@@ -295,8 +295,8 @@
context lattice_hom
begin
- notation le'.meet (infixl "\<sqinter>''" 50)
- notation le'.join (infixl "\<squnion>''" 50)
+ notation le'.meet (infixl \<open>\<sqinter>''\<close> 50)
+ notation le'.join (infixl \<open>\<squnion>''\<close> 50)
end
text \<open>The next example makes radical use of the short hand
--- 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>\<open>\\_\<close>")
+consts DUMMY :: 'a (\<open>\<^latex>\<open>\_\<close>\<close>)
(* THEOREMS *)
notation (Rule output)
- Pure.imp ("\<^latex>\<open>\\mbox{}\\inferrule{\\mbox{\<close>_\<^latex>\<open>}}\<close>\<^latex>\<open>{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+ Pure.imp (\<open>\<^latex>\<open>\mbox{}\inferrule{\mbox{\<close>_\<^latex>\<open>}}\<close>\<^latex>\<open>{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
syntax (Rule output)
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
- ("\<^latex>\<open>\\mbox{}\\inferrule{\<close>_\<^latex>\<open>}\<close>\<^latex>\<open>{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+ (\<open>\<^latex>\<open>\mbox{}\inferrule{\<close>_\<^latex>\<open>}\<close>\<^latex>\<open>{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
"_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms"
- ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\\\\\<close>/ _")
+ (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\\\<close>/ _\<close>)
- "_asm" :: "prop \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>")
+ "_asm" :: "prop \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close>)
notation (Axiom output)
- "Trueprop" ("\<^latex>\<open>\\mbox{}\\inferrule{\\mbox{}}{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+ "Trueprop" (\<open>\<^latex>\<open>\mbox{}\inferrule{\mbox{}}{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
notation (IfThen output)
- Pure.imp ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _/ \<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
+ Pure.imp (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _/ \<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
syntax (IfThen output)
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
- ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _ /\<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
- "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close> /\<^latex>\<open>{\\normalsize \\,\<close>and\<^latex>\<open>\\,}\<close>/ _")
- "_asm" :: "prop \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>")
+ (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _ /\<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
+ "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close> /\<^latex>\<open>{\normalsize \,\<close>and\<^latex>\<open>\,}\<close>/ _\<close>)
+ "_asm" :: "prop \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close>)
notation (IfThenNoBox output)
- Pure.imp ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _/ \<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
+ Pure.imp (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _/ \<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
syntax (IfThenNoBox output)
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
- ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _ /\<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
- "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^latex>\<open>{\\normalsize \\,\<close>and\<^latex>\<open>\\,}\<close>/ _")
- "_asm" :: "prop \<Rightarrow> asms" ("_")
+ (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _ /\<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
+ "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" (\<open>_ /\<^latex>\<open>{\normalsize \,\<close>and\<^latex>\<open>\,}\<close>/ _\<close>)
+ "_asm" :: "prop \<Rightarrow> asms" (\<open>_\<close>)
setup \<open>
Document_Output.antiquotation_pretty_source \<^binding>\<open>const_typ\<close>
--- 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.}
\<close>
(*<*)
-primrec valid :: "state \<Rightarrow> formula \<Rightarrow> bool" ("(_ \<Turnstile> _)" [80,80] 80) where
+primrec valid :: "state \<Rightarrow> formula \<Rightarrow> bool" (\<open>(_ \<Turnstile> _)\<close> [80,80] 80) where
"s \<Turnstile> Atom a = (a \<in> L s)" |
"s \<Turnstile> Neg f = (~(s \<Turnstile> f))" |
"s \<Turnstile> And f g = (s \<Turnstile> f \<and> s \<Turnstile> g)" |
--- 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{\<open>valid s f\<close>}. The definition is by recursion over the syntax:
\<close>
-primrec valid :: "state \<Rightarrow> formula \<Rightarrow> bool" ("(_ \<Turnstile> _)" [80,80] 80)
+primrec valid :: "state \<Rightarrow> formula \<Rightarrow> bool" (\<open>(_ \<Turnstile> _)\<close> [80,80] 80)
where
"s \<Turnstile> Atom a = (a \<in> L s)" |
"s \<Turnstile> Neg f = (\<not>(s \<Turnstile> f))" |
--- 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.
\<close>
-definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]" 60)
+definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl \<open>[+]\<close> 60)
where "A [+] B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
text \<open>
@@ -138,7 +138,7 @@
hide_const xor
setup \<open>Sign.add_path "version1"\<close>
(*>*)
-definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "\<oplus>" 60)
+definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl \<open>\<oplus>\<close> 60)
where "A \<oplus> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
(*<*)
setup \<open>Sign.local_path\<close>
@@ -156,10 +156,10 @@
hide_const xor
setup \<open>Sign.add_path "version2"\<close>
(*>*)
-definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "[+]\<ignore>" 60)
+definition xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl \<open>[+]\<ignore>\<close> 60)
where "A [+]\<ignore> B \<equiv> (A \<and> \<not> B) \<or> (\<not> A \<and> B)"
-notation (xsymbols) xor (infixl "\<oplus>\<ignore>" 60)
+notation (xsymbols) xor (infixl \<open>\<oplus>\<ignore>\<close> 60)
(*<*)
setup \<open>Sign.local_path\<close>
(*>*)
@@ -186,10 +186,10 @@
\<close>
datatype currency =
- Euro nat ("\<euro>")
- | Pounds nat ("\<pounds>")
- | Yen nat ("\<yen>")
- | Dollar nat ("$")
+ Euro nat (\<open>\<euro>\<close>)
+ | Pounds nat (\<open>\<pounds>\<close>)
+ | Yen nat (\<open>\<yen>\<close>)
+ | Dollar nat (\<open>$\<close>)
text \<open>
\noindent Here the mixfix annotations on the rightmost column happen
@@ -223,7 +223,7 @@
\<open>x \<approx> y\<close>. We assume that a constant \<open>sim\<close> of type
\<^typ>\<open>('a \<times> 'a) set\<close> has been introduced at this point.\<close>
(*<*)consts sim :: "('a \<times> 'a) set"(*>*)
-abbreviation sim2 :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<approx>" 50)
+abbreviation sim2 :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open>\<approx>\<close> 50)
where "x \<approx> y \<equiv> (x, y) \<in> sim"
text \<open>\noindent The given meta-equality is used as a rewrite rule
@@ -238,10 +238,10 @@
stems from Isabelle/HOL itself:
\<close>
-abbreviation not_equal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "~=\<ignore>" 50)
+abbreviation not_equal :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>~=\<ignore>\<close> 50)
where "x ~=\<ignore> y \<equiv> \<not> (x = y)"
-notation (xsymbols) not_equal (infix "\<noteq>\<ignore>" 50)
+notation (xsymbols) not_equal (infix \<open>\<noteq>\<ignore>\<close> 50)
text \<open>\noindent The notation \<open>\<noteq>\<close> is introduced separately to restrict it
to the \emph{xsymbols} mode.
--- 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 @@
\<close>
inductive_set
- rtc :: "('a \<times> 'a)set \<Rightarrow> ('a \<times> 'a)set" ("_*" [1000] 999)
+ rtc :: "('a \<times> 'a)set \<Rightarrow> ('a \<times> 'a)set" (\<open>_*\<close> [1000] 999)
for r :: "('a \<times> 'a)set"
where
rtc_refl[iff]: "(x,x) \<in> r*"
--- 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:
\<close>
-datatype nat = zero ("0") | Suc nat
+datatype nat = zero (\<open>0\<close>) | Suc nat
(*<*)
end
(*>*)
--- 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\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
nonterminal mtuple_args
syntax
- "" :: "'a \<Rightarrow> mtuple_args" ("_")
- "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" ("_,/ _")
- "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)")
+ "" :: "'a \<Rightarrow> mtuple_args" (\<open>_\<close>)
+ "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" (\<open>_,/ _\<close>)
+ "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" (\<open>(2\<lbrace>_,/ _\<rbrace>)\<close>)
syntax_consts
"_MTuple_args" "_MTuple" \<rightleftharpoons> MPair
translations
--- 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>\<open>Main\<close> as follows.
\<close>
-no_notation Nil ("[]") and Cons (infixr "#" 65) and append (infixr "@" 65)
+no_notation Nil (\<open>[]\<close>) and Cons (infixr \<open>#\<close> 65) and append (infixr \<open>@\<close> 65)
hide_type list
hide_const rev
-datatype 'a list = Nil ("[]")
- | Cons 'a "'a list" (infixr "#" 65)
+datatype 'a list = Nil (\<open>[]\<close>)
+ | Cons 'a "'a list" (infixr \<open>#\<close> 65)
text\<open>\noindent
The datatype\index{datatype@\isacommand {datatype} (command)}
@@ -47,7 +47,7 @@
in this order, because Isabelle insists on definition before use:
\<close>
-primrec app :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "@" 65) where
+primrec app :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr \<open>@\<close> 65) where
"[] @ ys = ys" |
"(x # xs) @ ys = x # (xs @ ys)"
--- 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 \<open>neutral\<close> together with its property:\<close>
class monoidl = semigroup +
- fixes neutral :: "'a" ("\<zero>")
+ fixes neutral :: "'a" (\<open>\<zero>\<close>)
assumes neutl: "\<zero> \<oplus> x = x"
text \<open>\noindent Again, we prove some instances, by providing
@@ -140,7 +140,7 @@
text \<open>\noindent To finish our small algebra example, we add a \<open>group\<close> class:\<close>
class group = monoidl +
- fixes inv :: "'a \<Rightarrow> 'a" ("\<div> _" [81] 80)
+ fixes inv :: "'a \<Rightarrow> 'a" (\<open>\<div> _\<close> [81] 80)
assumes invl: "\<div> x \<oplus> x = \<zero>"
text \<open>\noindent We continue with a further example for abstract
--- 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:\<close>
class plus =
- fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oplus>" 70)
+ fixes plus :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<oplus>\<close> 70)
text \<open>\noindent This introduces a new class @{class [source] plus},
along with a constant @{const [source] plus} with nice infix syntax.
--- 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 :: \<open>['a, 'a] \<Rightarrow> o\<close> (infixl \<open>\<noteq>\<close> 50)
where \<open>x \<noteq> y \<equiv> \<not> (x = y)\<close>
-syntax "_Uniq" :: "pttrn \<Rightarrow> o \<Rightarrow> o" ("(2\<exists>\<^sub>\<le>\<^sub>1 _./ _)" [0, 10] 10)
+syntax "_Uniq" :: "pttrn \<Rightarrow> o \<Rightarrow> o" (\<open>(2\<exists>\<^sub>\<le>\<^sub>1 _./ _)\<close> [0, 10] 10)
syntax_consts "_Uniq" \<rightleftharpoons> Uniq
translations "\<exists>\<^sub>\<le>\<^sub>1x. P" \<rightleftharpoons> "CONST Uniq (\<lambda>x. P)"
--- 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" (\<open>(3_ /= _ :/ _)\<close> [10,10,10] 5)
(*** Logical Connectives -- Type Formers ***)
- eq :: "['a,'a] => o" (infixl "=" 50)
+ eq :: "['a,'a] => o" (infixl \<open>=\<close> 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 \<open>&\<close> 35)
+ disj :: "[o,o] => o" (infixr \<open>|\<close> 30)
+ imp :: "[o,o] => o" (infixr \<open>-->\<close> 25)
(*Quantifiers*)
- All :: "('a => o) => o" (binder "ALL " 10)
- Ex :: "('a => o) => o" (binder "EX " 10)
+ All :: "('a => o) => o" (binder \<open>ALL \<close> 10)
+ Ex :: "('a => o) => o" (binder \<open>EX \<close> 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" (\<open>(1<_,/_>)\<close>)
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 \<open>lam \<close> 55)
+ App :: "[p,p]=>p" (infixl \<open>`\<close> 60)
+ alll :: "['a=>p]=>p" (binder \<open>all \<close> 55)
+ app :: "[p,'a]=>p" (infixl \<open>^\<close> 55)
+ exists :: "['a,p]=>p" (\<open>(1[_,/_])\<close>)
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" (\<open>(_ /: _)\<close> [51, 10] 5)
parse_translation \<open>
let fun proof_tr [p, P] = Syntax.const \<^const_syntax>\<open>Proof\<close> $ P $ p
@@ -155,14 +155,14 @@
(**** Definitions ****)
-definition Not :: "o => o" ("~ _" [40] 40)
+definition Not :: "o => o" (\<open>~ _\<close> [40] 40)
where not_def: "~P == P-->False"
-definition iff :: "[o,o] => o" (infixr "<->" 25)
+definition iff :: "[o,o] => o" (infixr \<open><->\<close> 25)
where "P<->Q == (P-->Q) & (Q-->P)"
(*Unique existence*)
-definition Ex1 :: "('a => o) => o" (binder "EX! " 10)
+definition Ex1 :: "('a => o) => o" (binder \<open>EX! \<close> 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*)
--- 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 (\<open>0\<close>) 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 \<open>+\<close> 60)
where "m + n == rec(m, n, %x y. Suc(y))"
--- 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 \<open>Hiding \<open><+>\<close> from \<^theory>\<open>HOL.Sum_Type\<close> until I come
up with better syntax here\<close>
-no_notation Sum_Type.Plus (infixr "<+>" 65)
+no_notation Sum_Type.Plus (infixr \<open><+>\<close> 65)
definition
- a_r_coset :: "[_, 'a set, 'a] \<Rightarrow> 'a set" (infixl "+>\<index>" 60)
+ a_r_coset :: "[_, 'a set, 'a] \<Rightarrow> 'a set" (infixl \<open>+>\<index>\<close> 60)
where "a_r_coset G = r_coset (add_monoid G)"
definition
- a_l_coset :: "[_, 'a, 'a set] \<Rightarrow> 'a set" (infixl "<+\<index>" 60)
+ a_l_coset :: "[_, 'a, 'a set] \<Rightarrow> 'a set" (infixl \<open><+\<index>\<close> 60)
where "a_l_coset G = l_coset (add_monoid G)"
definition
- A_RCOSETS :: "[_, 'a set] \<Rightarrow> ('a set)set" ("a'_rcosets\<index> _" [81] 80)
+ A_RCOSETS :: "[_, 'a set] \<Rightarrow> ('a set)set" (\<open>a'_rcosets\<index> _\<close> [81] 80)
where "A_RCOSETS G H = RCOSETS (add_monoid G) H"
definition
- set_add :: "[_, 'a set ,'a set] \<Rightarrow> 'a set" (infixl "<+>\<index>" 60)
+ set_add :: "[_, 'a set ,'a set] \<Rightarrow> 'a set" (infixl \<open><+>\<index>\<close> 60)
where "set_add G = set_mult (add_monoid G)"
definition
- A_SET_INV :: "[_,'a set] \<Rightarrow> 'a set" ("a'_set'_inv\<index> _" [81] 80)
+ A_SET_INV :: "[_,'a set] \<Rightarrow> 'a set" (\<open>a'_set'_inv\<index> _\<close> [81] 80)
where "A_SET_INV G H = SET_INV (add_monoid G) H"
definition
- a_r_congruent :: "[('a,'b)ring_scheme, 'a set] \<Rightarrow> ('a*'a)set" ("racong\<index>")
+ a_r_congruent :: "[('a,'b)ring_scheme, 'a set] \<Rightarrow> ('a*'a)set" (\<open>racong\<index>\<close>)
where "a_r_congruent G = r_congruent (add_monoid G)"
definition
- A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "A'_Mod" 65)
+ A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl \<open>A'_Mod\<close> 65)
\<comment> \<open>Actually defined for groups rather than monoids\<close>
where "A_FactGroup G H = FactGroup (add_monoid G) H"
--- 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 \<open>Definitions\<close>
-inductive iso_incl :: "'a ring \<Rightarrow> 'a ring \<Rightarrow> bool" (infixl "\<lesssim>" 65) for A B
+inductive iso_incl :: "'a ring \<Rightarrow> 'a ring \<Rightarrow> bool" (infixl \<open>\<lesssim>\<close> 65) for A B
where iso_inclI [intro]: "id \<in> ring_hom A B \<Longrightarrow> iso_incl A B"
definition law_restrict :: "('a, 'b) ring_scheme \<Rightarrow> 'a ring"
@@ -33,7 +33,7 @@
\<not> index_free \<P> (P, i) \<longrightarrow>
\<X>\<^bsub>(P, i)\<^esub> \<in> carrier L \<and> (ring.eval L) (\<sigma> P) \<X>\<^bsub>(P, i)\<^esub> = \<zero>\<^bsub>L\<^esub>) }"
-abbreviation (in ring) restrict_extensions :: "((('a list \<times> nat) multiset) \<Rightarrow> 'a) ring set" ("\<S>")
+abbreviation (in ring) restrict_extensions :: "((('a list \<times> nat) multiset) \<Rightarrow> 'a) ring set" (\<open>\<S>\<close>)
where "\<S> \<equiv> law_restrict ` extensions"
--- 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 ("_\<up>" [1000] 999)
-notation of_ac ("_\<down>" [1000] 999)
+notation to_ac (\<open>_\<up>\<close> [1000] 999)
+notation of_ac (\<open>_\<down>\<close> [1000] 999)
end
bundle alg_closure_syntax'
begin
-notation (output) to_ac ("_")
-notation (output) of_ac ("_")
+notation (output) to_ac (\<open>_\<close>)
+notation (output) of_ac (\<open>_\<close>)
end
--- 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 \<open>Structure with Carrier and Equivalence Relation \<open>eq\<close>\<close>
record 'a eq_object = "'a partial_object" +
- eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl ".=\<index>" 50)
+ eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>.=\<index>\<close> 50)
definition
- elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl ".\<in>\<index>" 50)
+ elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open>.\<in>\<index>\<close> 50)
where "x .\<in>\<^bsub>S\<^esub> A \<longleftrightarrow> (\<exists>y \<in> A. x .=\<^bsub>S\<^esub> y)"
definition
- set_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "{.=}\<index>" 50)
+ set_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open>{.=}\<index>\<close> 50)
where "A {.=}\<^bsub>S\<^esub> B \<longleftrightarrow> ((\<forall>x \<in> A. x .\<in>\<^bsub>S\<^esub> B) \<and> (\<forall>x \<in> B. x .\<in>\<^bsub>S\<^esub> A))"
definition
- eq_class_of :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set" ("class'_of\<index>")
+ eq_class_of :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set" (\<open>class'_of\<index>\<close>)
where "class_of\<^bsub>S\<^esub> x = {y \<in> carrier S. x .=\<^bsub>S\<^esub> y}"
definition
- eq_classes :: "_ \<Rightarrow> ('a set) set" ("classes\<index>")
+ eq_classes :: "_ \<Rightarrow> ('a set) set" (\<open>classes\<index>\<close>)
where "classes\<^bsub>S\<^esub> = {class_of\<^bsub>S\<^esub> x | x. x \<in> carrier S}"
definition
- eq_closure_of :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set" ("closure'_of\<index>")
+ eq_closure_of :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set" (\<open>closure'_of\<index>\<close>)
where "closure_of\<^bsub>S\<^esub> A = {y \<in> carrier S. y .\<in>\<^bsub>S\<^esub> A}"
definition
- eq_is_closed :: "_ \<Rightarrow> 'a set \<Rightarrow> bool" ("is'_closed\<index>")
+ eq_is_closed :: "_ \<Rightarrow> 'a set \<Rightarrow> bool" (\<open>is'_closed\<index>\<close>)
where "is_closed\<^bsub>S\<^esub> A \<longleftrightarrow> A \<subseteq> carrier S \<and> closure_of\<^bsub>S\<^esub> A = A"
abbreviation
- not_eq :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl ".\<noteq>\<index>" 50)
+ not_eq :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>.\<noteq>\<index>\<close> 50)
where "x .\<noteq>\<^bsub>S\<^esub> y \<equiv> \<not>(x .=\<^bsub>S\<^esub> y)"
abbreviation
- not_elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl ".\<notin>\<index>" 50)
+ not_elem :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open>.\<notin>\<index>\<close> 50)
where "x .\<notin>\<^bsub>S\<^esub> A \<equiv> \<not>(x .\<in>\<^bsub>S\<^esub> A)"
abbreviation
- set_not_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "{.\<noteq>}\<index>" 50)
+ set_not_eq :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl \<open>{.\<noteq>}\<index>\<close> 50)
where "A {.\<noteq>}\<^bsub>S\<^esub> B \<equiv> \<not>(A {.=}\<^bsub>S\<^esub> B)"
locale equivalence =
--- 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 \<open>Cosets and Quotient Groups\<close>
definition
- r_coset :: "[_, 'a set, 'a] \<Rightarrow> 'a set" (infixl "#>\<index>" 60)
+ r_coset :: "[_, 'a set, 'a] \<Rightarrow> 'a set" (infixl \<open>#>\<index>\<close> 60)
where "H #>\<^bsub>G\<^esub> a = (\<Union>h\<in>H. {h \<otimes>\<^bsub>G\<^esub> a})"
definition
- l_coset :: "[_, 'a, 'a set] \<Rightarrow> 'a set" (infixl "<#\<index>" 60)
+ l_coset :: "[_, 'a, 'a set] \<Rightarrow> 'a set" (infixl \<open><#\<index>\<close> 60)
where "a <#\<^bsub>G\<^esub> H = (\<Union>h\<in>H. {a \<otimes>\<^bsub>G\<^esub> h})"
definition
- RCOSETS :: "[_, 'a set] \<Rightarrow> ('a set)set" ("rcosets\<index> _" [81] 80)
+ RCOSETS :: "[_, 'a set] \<Rightarrow> ('a set)set" (\<open>rcosets\<index> _\<close> [81] 80)
where "rcosets\<^bsub>G\<^esub> H = (\<Union>a\<in>carrier G. {H #>\<^bsub>G\<^esub> a})"
definition
- set_mult :: "[_, 'a set ,'a set] \<Rightarrow> 'a set" (infixl "<#>\<index>" 60)
+ set_mult :: "[_, 'a set ,'a set] \<Rightarrow> 'a set" (infixl \<open><#>\<index>\<close> 60)
where "H <#>\<^bsub>G\<^esub> K = (\<Union>h\<in>H. \<Union>k\<in>K. {h \<otimes>\<^bsub>G\<^esub> k})"
definition
- SET_INV :: "[_,'a set] \<Rightarrow> 'a set" ("set'_inv\<index> _" [81] 80)
+ SET_INV :: "[_,'a set] \<Rightarrow> 'a set" (\<open>set'_inv\<index> _\<close> [81] 80)
where "set_inv\<^bsub>G\<^esub> H = (\<Union>h\<in>H. {inv\<^bsub>G\<^esub> h})"
@@ -35,7 +35,7 @@
assumes coset_eq: "(\<forall>x \<in> carrier G. H #> x = x <# H)"
abbreviation
- normal_rel :: "['a set, ('a, 'b) monoid_scheme] \<Rightarrow> bool" (infixl "\<lhd>" 60) where
+ normal_rel :: "['a set, ('a, 'b) monoid_scheme] \<Rightarrow> bool" (infixl \<open>\<lhd>\<close> 60) where
"H \<lhd> G \<equiv> normal H G"
lemma (in comm_group) subgroup_imp_normal: "subgroup A G \<Longrightarrow> A \<lhd> G"
@@ -659,7 +659,7 @@
subsubsection\<open>An Equivalence Relation\<close>
definition
- r_congruent :: "[('a,'b)monoid_scheme, 'a set] \<Rightarrow> ('a*'a)set" ("rcong\<index> _")
+ r_congruent :: "[('a,'b)monoid_scheme, 'a set] \<Rightarrow> ('a*'a)set" (\<open>rcong\<index> _\<close>)
where "rcong\<^bsub>G\<^esub> H = {(x,y). x \<in> carrier G \<and> y \<in> carrier G \<and> inv\<^bsub>G\<^esub> x \<otimes>\<^bsub>G\<^esub> y \<in> H}"
@@ -961,7 +961,7 @@
subsection \<open>Quotient Groups: Factorization of a Group\<close>
definition
- FactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "Mod" 65)
+ FactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl \<open>Mod\<close> 65)
\<comment> \<open>Actually defined for groups rather than monoids\<close>
where "FactGroup G H = \<lparr>carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
--- 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 \<open>Function definitions\<close>
-definition factor :: "[_, 'a, 'a] \<Rightarrow> bool" (infix "divides\<index>" 65)
+definition factor :: "[_, 'a, 'a] \<Rightarrow> bool" (infix \<open>divides\<index>\<close> 65)
where "a divides\<^bsub>G\<^esub> b \<longleftrightarrow> (\<exists>c\<in>carrier G. b = a \<otimes>\<^bsub>G\<^esub> c)"
-definition associated :: "[_, 'a, 'a] \<Rightarrow> bool" (infix "\<sim>\<index>" 55)
+definition associated :: "[_, 'a, 'a] \<Rightarrow> bool" (infix \<open>\<sim>\<index>\<close> 55)
where "a \<sim>\<^bsub>G\<^esub> b \<longleftrightarrow> a divides\<^bsub>G\<^esub> b \<and> b divides\<^bsub>G\<^esub> a"
abbreviation "division_rel G \<equiv> \<lparr>carrier = carrier G, eq = (\<sim>\<^bsub>G\<^esub>), le = (divides\<^bsub>G\<^esub>)\<rparr>"
@@ -821,7 +821,7 @@
definition wfactors ::"('a, _) monoid_scheme \<Rightarrow> 'a list \<Rightarrow> 'a \<Rightarrow> bool"
where "wfactors G fs a \<longleftrightarrow> (\<forall>x \<in> (set fs). irreducible G x) \<and> foldr (\<otimes>\<^bsub>G\<^esub>) fs \<one>\<^bsub>G\<^esub> \<sim>\<^bsub>G\<^esub> a"
-abbreviation list_assoc :: "('a, _) monoid_scheme \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "[\<sim>]\<index>" 44)
+abbreviation list_assoc :: "('a, _) monoid_scheme \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix \<open>[\<sim>]\<index>\<close> 44)
where "list_assoc G \<equiv> list_all2 (\<sim>\<^bsub>G\<^esub>)"
definition essentially_equal :: "('a, _) monoid_scheme \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
@@ -1993,11 +1993,11 @@
subsubsection \<open>Definitions\<close>
-definition isgcd :: "[('a,_) monoid_scheme, 'a, 'a, 'a] \<Rightarrow> bool" ("(_ gcdof\<index> _ _)" [81,81,81] 80)
+definition isgcd :: "[('a,_) monoid_scheme, 'a, 'a, 'a] \<Rightarrow> bool" (\<open>(_ gcdof\<index> _ _)\<close> [81,81,81] 80)
where "x gcdof\<^bsub>G\<^esub> a b \<longleftrightarrow> x divides\<^bsub>G\<^esub> a \<and> x divides\<^bsub>G\<^esub> b \<and>
(\<forall>y\<in>carrier G. (y divides\<^bsub>G\<^esub> a \<and> y divides\<^bsub>G\<^esub> b \<longrightarrow> y divides\<^bsub>G\<^esub> x))"
-definition islcm :: "[_, 'a, 'a, 'a] \<Rightarrow> bool" ("(_ lcmof\<index> _ _)" [81,81,81] 80)
+definition islcm :: "[_, 'a, 'a, 'a] \<Rightarrow> bool" (\<open>(_ lcmof\<index> _ _)\<close> [81,81,81] 80)
where "x lcmof\<^bsub>G\<^esub> a b \<longleftrightarrow> a divides\<^bsub>G\<^esub> x \<and> b divides\<^bsub>G\<^esub> x \<and>
(\<forall>y\<in>carrier G. (a divides\<^bsub>G\<^esub> y \<and> b divides\<^bsub>G\<^esub> y \<longrightarrow> x divides\<^bsub>G\<^esub> y))"
--- 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 \<Rightarrow> 'a list \<Rightarrow> bool"
where "dependent K U \<equiv> \<not> independent K U"
-definition over :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" (infixl "over" 65)
+definition over :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" (infixl \<open>over\<close> 65)
where "f over a = f a"
--- 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 \<Rightarrow> 'a) \<Rightarrow> 'a monoid list \<times> ('a \<Rightarrow> 'a) list \<Rightarrow> 'a monoid \<Rightarrow> 'a monoid list \<times> ('a \<Rightarrow> 'a) list"
- ("(3_ / \<longlongrightarrow>\<index> _)" [1000, 60])
+ (\<open>(3_ / \<longlongrightarrow>\<index> _)\<close> [1000, 60])
where "exact_seq_arrow f t G \<equiv> (G # (fst t), f # (snd t))"
--- 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 \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<cdot>" 70)
+ and f :: "'b \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<cdot>\<close> 70)
assumes left_commute:
"\<lbrakk>x \<in> B; y \<in> B; z \<in> D\<rbrakk> \<Longrightarrow> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)"
and f_closed [simp, intro!]: "!!x y. \<lbrakk>x \<in> B; y \<in> D\<rbrakk> \<Longrightarrow> f x y \<in> D"
@@ -246,7 +246,7 @@
locale ACeD =
fixes D :: "'a set"
- and f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<cdot>" 70)
+ and f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<cdot>\<close> 70)
and e :: 'a
assumes ident [simp]: "x \<in> D \<Longrightarrow> x \<cdot> e = x"
and commute: "\<lbrakk>x \<in> D; y \<in> D\<rbrakk> \<Longrightarrow> x \<cdot> y = y \<cdot> x"
@@ -307,7 +307,7 @@
syntax
"_finprod" :: "index \<Rightarrow> idt \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"
- ("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10)
+ (\<open>(3\<Otimes>__\<in>_. _)\<close> [1000, 0, 51, 10] 10)
syntax_consts
"_finprod" \<rightleftharpoons> finprod
translations
--- 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 \<open>Definition and basic properties\<close>
record ('a, 'b, 'c, 'd) galcon =
- orderA :: "('a, 'c) gorder_scheme" ("\<X>\<index>")
- orderB :: "('b, 'd) gorder_scheme" ("\<Y>\<index>")
- lower :: "'a \<Rightarrow> 'b" ("\<pi>\<^sup>*\<index>")
- upper :: "'b \<Rightarrow> 'a" ("\<pi>\<^sub>*\<index>")
+ orderA :: "('a, 'c) gorder_scheme" (\<open>\<X>\<index>\<close>)
+ orderB :: "('b, 'd) gorder_scheme" (\<open>\<Y>\<index>\<close>)
+ lower :: "'a \<Rightarrow> 'b" (\<open>\<pi>\<^sup>*\<index>\<close>)
+ upper :: "'b \<Rightarrow> 'a" (\<open>\<pi>\<^sub>*\<index>\<close>)
type_synonym ('a, 'b) galois = "('a, 'b, unit, unit) galcon"
abbreviation "inv_galcon G \<equiv> \<lparr> orderA = inv_gorder \<Y>\<^bsub>G\<^esub>, orderB = inv_gorder \<X>\<^bsub>G\<^esub>, lower = upper G, upper = lower G \<rparr>"
-definition comp_galcon :: "('b, 'c) galois \<Rightarrow> ('a, 'b) galois \<Rightarrow> ('a, 'c) galois" (infixr "\<circ>\<^sub>g" 85)
+definition comp_galcon :: "('b, 'c) galois \<Rightarrow> ('a, 'b) galois \<Rightarrow> ('a, 'c) galois" (infixr \<open>\<circ>\<^sub>g\<close> 85)
where "G \<circ>\<^sub>g F = \<lparr> orderA = orderA F, orderB = orderB G, lower = lower G \<circ> lower F, upper = upper F \<circ> upper G \<rparr>"
-definition id_galcon :: "'a gorder \<Rightarrow> ('a, 'a) galois" ("I\<^sub>g") where
+definition id_galcon :: "'a gorder \<Rightarrow> ('a, 'a) galois" (\<open>I\<^sub>g\<close>) where
"I\<^sub>g(A) = \<lparr> orderA = A, orderB = A, lower = id, upper = id \<rparr>"
--- 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 @@
\<close>
record 'a monoid = "'a partial_object" +
- mult :: "['a, 'a] \<Rightarrow> 'a" (infixl "\<otimes>\<index>" 70)
- one :: 'a ("\<one>\<index>")
+ mult :: "['a, 'a] \<Rightarrow> 'a" (infixl \<open>\<otimes>\<index>\<close> 70)
+ one :: 'a (\<open>\<one>\<index>\<close>)
definition
- m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" ("inv\<index> _" [81] 80)
+ m_inv :: "('a, 'b) monoid_scheme => 'a => 'a" (\<open>inv\<index> _\<close> [81] 80)
where "inv\<^bsub>G\<^esub> x = (THE y. y \<in> carrier G \<and> x \<otimes>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub> \<and> y \<otimes>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub>)"
definition
@@ -353,7 +353,7 @@
subsection \<open>Power\<close>
consts
- pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr "[^]\<index>" 75)
+ pow :: "[('a, 'm) monoid_scheme, 'a, 'b::semiring_1] => 'a" (infixr \<open>[^]\<index>\<close> 75)
overloading nat_pow == "pow :: [_, 'a, nat] => 'a"
begin
@@ -793,7 +793,7 @@
subsection \<open>Direct Products\<close>
definition
- DirProd :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<times> 'b) monoid" (infixr "\<times>\<times>" 80) where
+ DirProd :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<times> 'b) monoid" (infixr \<open>\<times>\<times>\<close> 80) where
"G \<times>\<times> H =
\<lparr>carrier = carrier G \<times> carrier H,
mult = (\<lambda>(g, h) (g', h'). (g \<otimes>\<^bsub>G\<^esub> g', h \<otimes>\<^bsub>H\<^esub> h')),
@@ -905,7 +905,7 @@
definition iso :: "_ => _ => ('a => 'b) set"
where "iso G H = {h. h \<in> hom G H \<and> bij_betw h (carrier G) (carrier H)}"
-definition is_iso :: "_ \<Rightarrow> _ \<Rightarrow> bool" (infixr "\<cong>" 60)
+definition is_iso :: "_ \<Rightarrow> _ \<Rightarrow> bool" (infixr \<open>\<cong>\<close> 60)
where "G \<cong> H = (iso G H \<noteq> {})"
definition mon where "mon G H = {f \<in> hom G H. inj_on f (carrier G)}"
--- 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) \<open>Ideals Generated by a Subset of \<^term>\<open>carrier R\<close>\<close>
-definition genideal :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set" ("Idl\<index> _" [80] 79)
+definition genideal :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a set" (\<open>Idl\<index> _\<close> [80] 79)
where "genideal R S = \<Inter>{I. ideal I R \<and> S \<subseteq> I}"
subsubsection \<open>Principal Ideals\<close>
@@ -411,7 +411,7 @@
text \<open>Generation of Principal Ideals in Commutative Rings\<close>
-definition cgenideal :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set" ("PIdl\<index> _" [80] 79)
+definition cgenideal :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set" (\<open>PIdl\<index> _\<close> [80] 79)
where "cgenideal R a = {x \<otimes>\<^bsub>R\<^esub> a | x. x \<in> carrier R}"
text \<open>genhideal (?) really generates an ideal\<close>
--- 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 \<open>In this section, we study the structure of the set of ideals of a given ring.\<close>
inductive_set
- ideal_prod :: "[ ('a, 'b) ring_scheme, 'a set, 'a set ] \<Rightarrow> 'a set" (infixl "\<cdot>\<index>" 80)
+ ideal_prod :: "[ ('a, 'b) ring_scheme, 'a set, 'a set ] \<Rightarrow> 'a set" (infixl \<open>\<cdot>\<index>\<close> 80)
for R and I and J (* both I and J are supposed ideals *) where
prod: "\<lbrakk> i \<in> I; j \<in> J \<rbrakk> \<Longrightarrow> i \<otimes>\<^bsub>R\<^esub> j \<in> ideal_prod R I J"
| sum: "\<lbrakk> s1 \<in> ideal_prod R I J; s2 \<in> ideal_prod R I J \<rbrakk> \<Longrightarrow> s1 \<oplus>\<^bsub>R\<^esub> s2 \<in> ideal_prod R I J"
--- 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 \<Rightarrow> ('c multiset \<Rightarrow> 'a)"
where "indexed_const k = (\<lambda>m. if m = {#} then k else \<zero>)"
-definition (in ring) indexed_pmult :: "('c multiset \<Rightarrow> 'a) \<Rightarrow> 'c \<Rightarrow> ('c multiset \<Rightarrow> 'a)" (infixl "\<Otimes>" 65)
+definition (in ring) indexed_pmult :: "('c multiset \<Rightarrow> 'a) \<Rightarrow> 'c \<Rightarrow> ('c multiset \<Rightarrow> 'a)" (infixl \<open>\<Otimes>\<close> 65)
where "indexed_pmult P i = (\<lambda>m. if i \<in># m then P (m - {# i #}) else \<zero>)"
-definition (in ring) indexed_padd :: "_ \<Rightarrow> _ \<Rightarrow> ('c multiset \<Rightarrow> 'a)" (infixl "\<Oplus>" 65)
+definition (in ring) indexed_padd :: "_ \<Rightarrow> _ \<Rightarrow> ('c multiset \<Rightarrow> 'a)" (infixl \<open>\<Oplus>\<close> 65)
where "indexed_padd P Q = (\<lambda>m. (P m) \<oplus> (Q m))"
-definition (in ring) indexed_var :: "'c \<Rightarrow> ('c multiset \<Rightarrow> 'a)" ("\<X>\<index>")
+definition (in ring) indexed_var :: "'c \<Rightarrow> ('c multiset \<Rightarrow> 'a)" (\<open>\<X>\<index>\<close>)
where "indexed_var i = (indexed_const \<one>) \<Otimes> i"
definition (in ring) index_free :: "('c multiset \<Rightarrow> 'a) \<Rightarrow> 'c \<Rightarrow> bool"
@@ -38,7 +38,7 @@
definition (in ring) carrier_coeff :: "('c multiset \<Rightarrow> 'a) \<Rightarrow> bool"
where "carrier_coeff P \<longleftrightarrow> (\<forall>m. P m \<in> carrier R)"
-inductive_set (in ring) indexed_pset :: "'c set \<Rightarrow> 'a set \<Rightarrow> ('c multiset \<Rightarrow> 'a) set" ("_ [\<X>\<index>]" 80)
+inductive_set (in ring) indexed_pset :: "'c set \<Rightarrow> 'a set \<Rightarrow> ('c multiset \<Rightarrow> 'a) set" (\<open>_ [\<X>\<index>]\<close> 80)
for I and K where
indexed_const: "k \<in> K \<Longrightarrow> indexed_const k \<in> (K[\<X>\<^bsub>I\<^esub>])"
| indexed_padd: "\<lbrakk> P \<in> (K[\<X>\<^bsub>I\<^esub>]); Q \<in> (K[\<X>\<^bsub>I\<^esub>]) \<rbrakk> \<Longrightarrow> P \<Oplus> Q \<in> (K[\<X>\<^bsub>I\<^esub>])"
--- 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 \<open>\<open>\<Z>\<close>: The Set of Integers as Algebraic Structure\<close>
-abbreviation int_ring :: "int ring" ("\<Z>")
+abbreviation int_ring :: "int ring" (\<open>\<Z>\<close>)
where "int_ring \<equiv> \<lparr>carrier = UNIV, mult = (*), one = 1, zero = 0, add = (+)\<rparr>"
lemma int_Zcarr [intro!, simp]: "k \<in> carrier \<Z>"
--- 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 \<open>Supremum and infimum\<close>
definition
- sup :: "[_, 'a set] => 'a" ("\<Squnion>\<index>_" [90] 90)
+ sup :: "[_, 'a set] => 'a" (\<open>\<Squnion>\<index>_\<close> [90] 90)
where "\<Squnion>\<^bsub>L\<^esub>A = (SOME x. least L x (Upper L A))"
definition
- inf :: "[_, 'a set] => 'a" ("\<Sqinter>\<index>_" [90] 90)
+ inf :: "[_, 'a set] => 'a" (\<open>\<Sqinter>\<index>_\<close> [90] 90)
where "\<Sqinter>\<^bsub>L\<^esub>A = (SOME x. greatest L x (Lower L A))"
definition supr ::
@@ -31,10 +31,10 @@
where "infi L A f = \<Sqinter>\<^bsub>L\<^esub>(f ` A)"
syntax
- "_inf1" :: "('a, 'b) gorder_scheme \<Rightarrow> pttrns \<Rightarrow> 'a \<Rightarrow> 'a" ("(3IINF\<index> _./ _)" [0, 10] 10)
- "_inf" :: "('a, 'b) gorder_scheme \<Rightarrow> pttrn \<Rightarrow> 'c set \<Rightarrow> 'a \<Rightarrow> 'a" ("(3IINF\<index> _:_./ _)" [0, 0, 10] 10)
- "_sup1" :: "('a, 'b) gorder_scheme \<Rightarrow> pttrns \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SSUP\<index> _./ _)" [0, 10] 10)
- "_sup" :: "('a, 'b) gorder_scheme \<Rightarrow> pttrn \<Rightarrow> 'c set \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SSUP\<index> _:_./ _)" [0, 0, 10] 10)
+ "_inf1" :: "('a, 'b) gorder_scheme \<Rightarrow> pttrns \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(3IINF\<index> _./ _)\<close> [0, 10] 10)
+ "_inf" :: "('a, 'b) gorder_scheme \<Rightarrow> pttrn \<Rightarrow> 'c set \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(3IINF\<index> _:_./ _)\<close> [0, 0, 10] 10)
+ "_sup1" :: "('a, 'b) gorder_scheme \<Rightarrow> pttrns \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(3SSUP\<index> _./ _)\<close> [0, 10] 10)
+ "_sup" :: "('a, 'b) gorder_scheme \<Rightarrow> pttrn \<Rightarrow> 'c set \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(3SSUP\<index> _:_./ _)\<close> [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 "\<squnion>\<index>" 65)
+ join :: "[_, 'a, 'a] => 'a" (infixl \<open>\<squnion>\<index>\<close> 65)
where "x \<squnion>\<^bsub>L\<^esub> y = \<Squnion>\<^bsub>L\<^esub>{x, y}"
definition
- meet :: "[_, 'a, 'a] => 'a" (infixl "\<sqinter>\<index>" 70)
+ meet :: "[_, 'a, 'a] => 'a" (infixl \<open>\<sqinter>\<index>\<close> 70)
where "x \<sqinter>\<^bsub>L\<^esub> y = \<Sqinter>\<^bsub>L\<^esub>{x, y}"
definition
- LEAST_FP :: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a" ("LFP\<index>") where
+ LEAST_FP :: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a" (\<open>LFP\<index>\<close>) where
"LEAST_FP L f = \<Sqinter>\<^bsub>L\<^esub> {u \<in> carrier L. f u \<sqsubseteq>\<^bsub>L\<^esub> u}" \<comment> \<open>least fixed point\<close>
definition
- GREATEST_FP:: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a" ("GFP\<index>") where
+ GREATEST_FP:: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a" (\<open>GFP\<index>\<close>) where
"GREATEST_FP L f = \<Squnion>\<^bsub>L\<^esub> {u \<in> carrier L. u \<sqsubseteq>\<^bsub>L\<^esub> f u}" \<comment> \<open>greatest fixed point\<close>
--- 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] \<Rightarrow> ('a set)set" ("lcosets\<index> _" [81] 80)
+ LCOSETS :: "[_, 'a set] \<Rightarrow> ('a set)set" (\<open>lcosets\<index> _\<close> [81] 80)
where "lcosets\<^bsub>G\<^esub> H = (\<Union>a\<in>carrier G. {a <#\<^bsub>G\<^esub> H})"
definition
- LFactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "LMod" 65)
+ LFactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl \<open>LMod\<close> 65)
\<comment> \<open>Actually defined for groups rather than monoids\<close>
where "LFactGroup G H = \<lparr>carrier = lcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
--- 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 \<open>Definitions\<close>
record ('a, 'b) module = "'b ring" +
- smult :: "['a, 'b] => 'b" (infixl "\<odot>\<index>" 70)
+ smult :: "['a, 'b] => 'b" (infixl \<open>\<odot>\<index>\<close> 70)
locale module = R?: cring + M?: abelian_group M for M (structure) +
assumes smult_closed [simp, intro]:
--- 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 \<le> x \<and> x \<le> m \<and> coprime x m}"
notation (latex output)
- phi' ("\<phi> _")
+ phi' (\<open>\<phi> _\<close>)
lemma phi'_nonzero:
assumes "m > 0"
--- 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 \<open>Partial Orders\<close>
record 'a gorder = "'a eq_object" +
- le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>\<index>" 50)
+ le :: "['a, 'a] => bool" (infixl \<open>\<sqsubseteq>\<index>\<close> 50)
abbreviation inv_gorder :: "_ \<Rightarrow> 'a gorder" where
"inv_gorder L \<equiv>
@@ -40,7 +40,7 @@
x \<sqsubseteq> z \<longleftrightarrow> y \<sqsubseteq> w"
definition
- lless :: "[_, 'a, 'a] => bool" (infixl "\<sqsubset>\<index>" 50)
+ lless :: "[_, 'a, 'a] => bool" (infixl \<open>\<sqsubset>\<index>\<close> 50)
where "x \<sqsubset>\<^bsub>L\<^esub> y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> y \<and> x .\<noteq>\<^bsub>L\<^esub> y"
subsubsection \<open>The order relation\<close>
@@ -411,7 +411,7 @@
subsubsection \<open>Intervals\<close>
definition
- at_least_at_most :: "('a, 'c) gorder_scheme \<Rightarrow> 'a => 'a => 'a set" ("(1\<lbrace>_.._\<rbrace>\<index>)")
+ at_least_at_most :: "('a, 'c) gorder_scheme \<Rightarrow> 'a => 'a => 'a set" (\<open>(1\<lbrace>_.._\<rbrace>\<index>)\<close>)
where "\<lbrace>l..u\<rbrace>\<^bsub>A\<^esub> = {x \<in> carrier A. l \<sqsubseteq>\<^bsub>A\<^esub> x \<and> x \<sqsubseteq>\<^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 \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" ("Mono\<index>")
+abbreviation Monotone :: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" (\<open>Mono\<index>\<close>)
where "Monotone L f \<equiv> isotone L L f"
lemma use_iso1:
@@ -478,7 +478,7 @@
subsubsection \<open>Idempotent functions\<close>
definition idempotent ::
- "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" ("Idem\<index>") where
+ "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" (\<open>Idem\<index>\<close>) where
"idempotent L f \<equiv> \<forall>x\<in>carrier L. f (f x) .=\<^bsub>L\<^esub> f x"
lemma (in weak_partial_order) idempotent:
@@ -568,11 +568,11 @@
subsection \<open>Bounded Orders\<close>
definition
- top :: "_ => 'a" ("\<top>\<index>") where
+ top :: "_ => 'a" (\<open>\<top>\<index>\<close>) where
"\<top>\<^bsub>L\<^esub> = (SOME x. greatest L x (carrier L))"
definition
- bottom :: "_ => 'a" ("\<bottom>\<index>") where
+ bottom :: "_ => 'a" (\<open>\<bottom>\<index>\<close>) where
"\<bottom>\<^bsub>L\<^esub> = (SOME x. least L x (carrier L))"
locale weak_partial_order_bottom = weak_partial_order L for L (structure) +
--- 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 :: "_ \<Rightarrow> ('a list) ring"
where "poly_ring R \<equiv> univ_poly R (carrier R)"
-abbreviation pirreducible :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" ("pirreducible\<index>")
+abbreviation pirreducible :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" (\<open>pirreducible\<index>\<close>)
where "pirreducible\<^bsub>R\<^esub> K p \<equiv> ring_irreducible\<^bsub>(univ_poly R K)\<^esub> p"
-abbreviation pprime :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" ("pprime\<index>")
+abbreviation pprime :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" (\<open>pprime\<index>\<close>)
where "pprime\<^bsub>R\<^esub> K p \<equiv> ring_prime\<^bsub>(univ_poly R K)\<^esub> p"
-definition pdivides :: "_ \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "pdivides\<index>" 65)
+definition pdivides :: "_ \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix \<open>pdivides\<index>\<close> 65)
where "p pdivides\<^bsub>R\<^esub> q = p divides\<^bsub>(univ_poly R (carrier R))\<^esub> q"
-definition rupture :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> (('a list) set) ring" ("Rupt\<index>")
+definition rupture :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> (('a list) set) ring" (\<open>Rupt\<index>\<close>)
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 \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> ('a list) set"
@@ -309,10 +309,10 @@
definition (in ring) long_division :: "'a list \<Rightarrow> 'a list \<Rightarrow> ('a list \<times> 'a list)"
where "long_division p q = (THE t. long_divides p q t)"
-definition (in ring) pdiv :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "pdiv" 65)
+definition (in ring) pdiv :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl \<open>pdiv\<close> 65)
where "p pdiv q = (if q = [] then [] else fst (long_division p q))"
-definition (in ring) pmod :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "pmod" 65)
+definition (in ring) pmod :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl \<open>pmod\<close> 65)
where "p pmod q = (if q = [] then p else snd (long_division p q))"
--- 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 \<Rightarrow> nat"
where "degree p \<equiv> length p - 1"
-definition polynomial :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" ("polynomial\<index>")
+definition polynomial :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" (\<open>polynomial\<index>\<close>)
where "polynomial\<^bsub>R\<^esub> K p \<longleftrightarrow> p = [] \<or> (set p \<subseteq> K \<and> lead_coeff p \<noteq> \<zero>\<^bsub>R\<^esub>)"
definition (in ring) monom :: "'a \<Rightarrow> nat \<Rightarrow> 'a list"
@@ -1359,7 +1359,7 @@
subsection \<open>Algebraic Structure of Polynomials\<close>
-definition univ_poly :: "('a, 'b) ring_scheme \<Rightarrow>'a set \<Rightarrow> ('a list) ring" ("_ [X]\<index>" 80)
+definition univ_poly :: "('a, 'b) ring_scheme \<Rightarrow>'a set \<Rightarrow> ('a list) ring" (\<open>_ [X]\<index>\<close> 80)
where "univ_poly R K =
\<lparr> carrier = { p. polynomial\<^bsub>R\<^esub> K p },
mult = ring.poly_mult R,
@@ -2111,7 +2111,7 @@
subsection \<open>The X Variable\<close>
-definition var :: "_ \<Rightarrow> 'a list" ("X\<index>")
+definition var :: "_ \<Rightarrow> 'a list" (\<open>X\<index>\<close>)
where "X\<^bsub>R\<^esub> = [ \<one>\<^bsub>R\<^esub>, \<zero>\<^bsub>R\<^esub> ]"
lemma (in ring) eval_var:
--- 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 \<open>Multiplication on Cosets\<close>
definition rcoset_mult :: "[('a, _) ring_scheme, 'a set, 'a set, 'a set] \<Rightarrow> 'a set"
- ("[mod _:] _ \<Otimes>\<index> _" [81,81,81] 80)
+ (\<open>[mod _:] _ \<Otimes>\<index> _\<close> [81,81,81] 80)
where "rcoset_mult R I A B = (\<Union>a\<in>A. \<Union>b\<in>B. I +>\<^bsub>R\<^esub> (a \<otimes>\<^bsub>R\<^esub> b))"
@@ -47,7 +47,7 @@
subsection \<open>Quotient Ring Definition\<close>
definition FactRing :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) ring"
- (infixl "Quot" 65)
+ (infixl \<open>Quot\<close> 65)
where "FactRing R I =
\<lparr>carrier = a_rcosets\<^bsub>R\<^esub> I, mult = rcoset_mult R I,
one = (I +>\<^bsub>R\<^esub> \<one>\<^bsub>R\<^esub>), zero = I, add = set_add R\<rparr>"
@@ -539,7 +539,7 @@
where "ring_iso R S = { h. h \<in> ring_hom R S \<and> bij_betw h (carrier R) (carrier S) }"
definition
- is_ring_iso :: "_ \<Rightarrow> _ \<Rightarrow> bool" (infixr "\<simeq>" 60)
+ is_ring_iso :: "_ \<Rightarrow> _ \<Rightarrow> bool" (infixr \<open>\<simeq>\<close> 60)
where "R \<simeq> S = (ring_iso R S \<noteq> {})"
definition
--- 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 \<open>Abelian Groups\<close>
record 'a ring = "'a monoid" +
- zero :: 'a ("\<zero>\<index>")
- add :: "['a, 'a] \<Rightarrow> 'a" (infixl "\<oplus>\<index>" 65)
+ zero :: 'a (\<open>\<zero>\<index>\<close>)
+ add :: "['a, 'a] \<Rightarrow> 'a" (infixl \<open>\<oplus>\<index>\<close> 65)
abbreviation
add_monoid :: "('a, 'm) ring_scheme \<Rightarrow> ('a, 'm) monoid_scheme"
@@ -23,15 +23,15 @@
text \<open>Derived operations.\<close>
definition
- a_inv :: "[('a, 'm) ring_scheme, 'a ] \<Rightarrow> 'a" ("\<ominus>\<index> _" [81] 80)
+ a_inv :: "[('a, 'm) ring_scheme, 'a ] \<Rightarrow> 'a" (\<open>\<ominus>\<index> _\<close> [81] 80)
where "a_inv R = m_inv (add_monoid R)"
definition
- a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" ("(_ \<ominus>\<index> _)" [65,66] 65)
+ a_minus :: "[('a, 'm) ring_scheme, 'a, 'a] => 'a" (\<open>(_ \<ominus>\<index> _)\<close> [65,66] 65)
where "x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"
definition
- add_pow :: "[_, ('b :: semiring_1), 'a] \<Rightarrow> 'a" ("[_] \<cdot>\<index> _" [81, 81] 80)
+ add_pow :: "[_, ('b :: semiring_1), 'a] \<Rightarrow> 'a" (\<open>[_] \<cdot>\<index> _\<close> [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 \<Rightarrow> idt \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b"
- ("(3\<Oplus>__\<in>_. _)" [1000, 0, 51, 10] 10)
+ (\<open>(3\<Oplus>__\<in>_. _)\<close> [1000, 0, 51, 10] 10)
syntax_consts
"_finsum" \<rightleftharpoons> finsum
translations
--- 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 @@
" \<lbrakk> a \<in> carrier R - { \<zero> }; b \<in> carrier R - { \<zero> } \<rbrakk> \<Longrightarrow>
\<exists>q r. q \<in> carrier R \<and> r \<in> carrier R \<and> a = (b \<otimes> q) \<oplus> r \<and> ((r = \<zero>) \<or> (\<phi> r < \<phi> b))"
-definition ring_irreducible :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> bool" ("ring'_irreducible\<index>")
+definition ring_irreducible :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> bool" (\<open>ring'_irreducible\<index>\<close>)
where "ring_irreducible\<^bsub>R\<^esub> a \<longleftrightarrow> (a \<noteq> \<zero>\<^bsub>R\<^esub>) \<and> (irreducible R a)"
-definition ring_prime :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> bool" ("ring'_prime\<index>")
+definition ring_prime :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> bool" (\<open>ring'_prime\<index>\<close>)
where "ring_prime\<^bsub>R\<^esub> a \<longleftrightarrow> (a \<noteq> \<zero>\<^bsub>R\<^esub>) \<and> (prime R a)"
--- 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\<open>Derived set (set of limit points)\<close>
-definition derived_set_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "derived'_set'_of" 80)
+definition derived_set_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl \<open>derived'_set'_of\<close> 80)
where "X derived_set_of S \<equiv>
{x \<in> topspace X.
(\<forall>T. x \<in> T \<and> openin X T \<longrightarrow> (\<exists>y\<noteq>x. y \<in> S \<and> y \<in> T))}"
@@ -715,7 +715,7 @@
subsection\<open> Closure with respect to a topological space\<close>
-definition closure_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr "closure'_of" 80)
+definition closure_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr \<open>closure'_of\<close> 80)
where "X closure_of S \<equiv> {x \<in> topspace X. \<forall>T. x \<in> T \<and> openin X T \<longrightarrow> (\<exists>y \<in> S. y \<in> T)}"
lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \<inter> S)"
@@ -947,7 +947,7 @@
text\<open> Interior with respect to a topological space. \<close>
-definition interior_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr "interior'_of" 80)
+definition interior_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr \<open>interior'_of\<close> 80)
where "X interior_of S \<equiv> {x. \<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> S}"
lemma interior_of_restrict:
@@ -1113,7 +1113,7 @@
subsection \<open>Frontier with respect to topological space \<close>
-definition frontier_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr "frontier'_of" 80)
+definition frontier_of :: "'a topology \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixr \<open>frontier'_of\<close> 80)
where "X frontier_of S \<equiv> X closure_of S - X interior_of S"
lemma frontier_of_closures:
@@ -3048,7 +3048,7 @@
subsection\<open>Relation of homeomorphism between topological spaces\<close>
-definition homeomorphic_space (infixr "homeomorphic'_space" 50)
+definition homeomorphic_space (infixr \<open>homeomorphic'_space\<close> 50)
where "X homeomorphic_space Y \<equiv> \<exists>f g. homeomorphic_maps X Y f g"
lemma homeomorphic_space_refl [iff]: "X homeomorphic_space X"
--- 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>\<open>tag unimportant\<close> \<open>A function constant on a set\<close>
-definition constant_on (infixl "(constant'_on)" 50)
+definition constant_on (infixl \<open>(constant'_on)\<close> 50)
where "f constant_on A \<equiv> \<exists>y. \<forall>x\<in>A. f x = y"
lemma constant_on_subset: "\<lbrakk>f constant_on A; B \<subseteq> A\<rbrakk> \<Longrightarrow> f constant_on B"
@@ -892,7 +892,7 @@
where "retraction S T r \<longleftrightarrow>
T \<subseteq> S \<and> continuous_on S r \<and> r \<in> S \<rightarrow> T \<and> (\<forall>x\<in>T. r x = x)"
-definition\<^marker>\<open>tag important\<close> retract_of (infixl "retract'_of" 50) where
+definition\<^marker>\<open>tag important\<close> retract_of (infixl \<open>retract'_of\<close> 50) where
"T retract_of S \<longleftrightarrow> (\<exists>r. retraction S T r)"
lemma retraction_idempotent: "retraction S T r \<Longrightarrow> x \<in> S \<Longrightarrow> r (r x) = r x"
@@ -1111,7 +1111,7 @@
subsection\<open>Retractions on a topological space\<close>
-definition retract_of_space :: "'a set \<Rightarrow> 'a topology \<Rightarrow> bool" (infix "retract'_of'_space" 50)
+definition retract_of_space :: "'a set \<Rightarrow> 'a topology \<Rightarrow> bool" (infix \<open>retract'_of'_space\<close> 50)
where "S retract_of_space X
\<equiv> S \<subseteq> topspace X \<and> (\<exists>r. continuous_map X (subtopology X S) r \<and> (\<forall>x \<in> S. r x = x))"
--- 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>\<open>tag important\<close> pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
+definition\<^marker>\<open>tag important\<close> pair_measure (infixr \<open>\<Otimes>\<^sub>M\<close> 80) where
"A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B)
{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
(\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
--- 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>\<open>tag important\<close> lebesgue_integral ("integral\<^sup>L") where
+definition\<^marker>\<open>tag important\<close> lebesgue_integral (\<open>integral\<^sup>L\<close>) where
"integral\<^sup>L M f = (if \<exists>x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)"
syntax
- "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> 'a measure \<Rightarrow> real" ("\<integral>((2 _./ _)/ \<partial>_)" [60,61] 110)
+ "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> 'a measure \<Rightarrow> real" (\<open>\<integral>((2 _./ _)/ \<partial>_)\<close> [60,61] 110)
syntax_consts
"_lebesgue_integral" == lebesgue_integral
@@ -899,7 +899,7 @@
"\<integral> x. f \<partial>M" == "CONST lebesgue_integral M (\<lambda>x. f)"
syntax
- "_ascii_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60)
+ "_ascii_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" (\<open>(3LINT (1_)/|(_)./ _)\<close> [0,110,60] 60)
syntax_consts
"_ascii_lebesgue_integral" == lebesgue_integral
--- 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 "<e" 50)
+ eucl_less (infix \<open><e\<close> 50)
lemma box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
and box_co: "{x. a \<le> x \<and> x <e b} = {a..} \<inter> {x. x <e b}"
@@ -1854,7 +1854,7 @@
unfolding Measurable.pred_def by (rule borel_measurable_less)
no_notation
- eucl_less (infix "<e" 50)
+ eucl_less (infix \<open><e\<close> 50)
lemma borel_measurable_Max2[measurable (raw)]:
fixes f::"_ \<Rightarrow> _ \<Rightarrow> 'a::{second_countable_topology, dense_linorder, linorder_topology}"
--- 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>\<open>tag important\<close> "bcontfun = {f. continuous_on UNIV f \<and> bounded (range f)}"
-typedef (overloaded) ('a, 'b) bcontfun ("(_ \<Rightarrow>\<^sub>C /_)" [22, 21] 21) =
+typedef (overloaded) ('a, 'b) bcontfun (\<open>(_ \<Rightarrow>\<^sub>C /_)\<close> [22, 21] 21) =
"bcontfun::('a::topological_space \<Rightarrow> 'b::metric_space) set"
morphisms apply_bcontfun Bcontfun
by (auto intro: continuous_intros simp: bounded_def bcontfun_def)
--- 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 \<open>Type of bounded linear functions\<close>
-typedef\<^marker>\<open>tag important\<close> (overloaded) ('a, 'b) blinfun ("(_ \<Rightarrow>\<^sub>L /_)" [22, 21] 21) =
+typedef\<^marker>\<open>tag important\<close> (overloaded) ('a, 'b) blinfun (\<open>(_ \<Rightarrow>\<^sub>L /_)\<close> [22, 21] 21) =
"{f::'a::real_normed_vector\<Rightarrow>'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 \<Rightarrow>\<^sub>L 'b::real_normed_vector \<Rightarrow>
'c::real_normed_vector \<Rightarrow>\<^sub>L 'a \<Rightarrow>
- 'c \<Rightarrow>\<^sub>L 'b" (infixl "o\<^sub>L" 55) is "(o)"
+ 'c \<Rightarrow>\<^sub>L 'b" (infixl \<open>o\<^sub>L\<close> 55) is "(o)"
parametric comp_transfer
unfolding o_def
by (rule bounded_linear_compose)
--- 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 \<open>*b\<close> 75)
+ and scaleC :: "('a => 'c::ab_group_add => 'c)" (infixr \<open>*c\<close> 75)
and BasisB :: "('b set)"
and f :: "('b=>'c)"
--- 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\<open>Holomorphic functions\<close>
definition\<^marker>\<open>tag important\<close> holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool"
- (infixl "(holomorphic'_on)" 50)
+ (infixl \<open>(holomorphic'_on)\<close> 50)
where "f holomorphic_on s \<equiv> \<forall>x\<in>s. f field_differentiable (at x within s)"
named_theorems\<^marker>\<open>tag important\<close> holomorphic_intros "structural introduction rules for holomorphic_on"
@@ -375,7 +375,7 @@
subsection\<open>Analyticity on a set\<close>
-definition\<^marker>\<open>tag important\<close> analytic_on (infixl "(analytic'_on)" 50)
+definition\<^marker>\<open>tag important\<close> analytic_on (infixl \<open>(analytic'_on)\<close> 50)
where "f analytic_on S \<equiv> \<forall>x \<in> S. \<exists>\<epsilon>. 0 < \<epsilon> \<and> f holomorphic_on (ball x \<epsilon>)"
named_theorems\<^marker>\<open>tag important\<close> analytic_intros "introduction rules for proving analyticity"
--- 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 \<comment>\<open>locally disable syntax for set product, to avoid warnings\<close>
-definition\<^marker>\<open>tag important\<close> cross3 :: "[real^3, real^3] \<Rightarrow> real^3" (infixr "\<times>" 80)
+definition\<^marker>\<open>tag important\<close> cross3 :: "[real^3, real^3] \<Rightarrow> real^3" (infixr \<open>\<times>\<close> 80)
where "a \<times> b \<equiv>
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 "\<times>" 80)
-no_notation Product_Type.Times (infixr "\<times>" 80)
+notation cross3 (infixr \<open>\<times>\<close> 80)
+no_notation Product_Type.Times (infixr \<open>\<times>\<close> 80)
end
bundle no_cross3_syntax begin
-no_notation cross3 (infixr "\<times>" 80)
-notation Product_Type.Times (infixr "\<times>" 80)
+no_notation cross3 (infixr \<open>\<times>\<close> 80)
+notation Product_Type.Times (infixr \<open>\<times>\<close> 80)
end
unbundle cross3_syntax
--- 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>\<open>tag important\<close>
differentiable_on :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool"
- (infix "differentiable'_on" 50)
+ (infix \<open>differentiable'_on\<close> 50)
where "f differentiable_on s \<longleftrightarrow> (\<forall>x\<in>s. f differentiable (at x within s))"
lemma differentiableI: "(f has_derivative f') net \<Longrightarrow> f differentiable net"
@@ -2008,7 +2008,7 @@
subsection \<open>Field differentiability\<close>
definition\<^marker>\<open>tag important\<close> field_differentiable :: "['a \<Rightarrow> 'a::real_normed_field, 'a filter] \<Rightarrow> bool"
- (infixr "(field'_differentiable)" 50)
+ (infixr \<open>(field'_differentiable)\<close> 50)
where "f field_differentiable F \<equiv> \<exists>f'. (f has_field_derivative f') F"
lemma field_differentiable_imp_differentiable:
@@ -3089,7 +3089,7 @@
subsection\<^marker>\<open>tag unimportant\<close> \<open>Piecewise differentiable functions\<close>
definition piecewise_differentiable_on
- (infixr "piecewise'_differentiable'_on" 50)
+ (infixr \<open>piecewise'_differentiable'_on\<close> 50)
where "f piecewise_differentiable_on i \<equiv>
continuous_on i f \<and>
(\<exists>S. finite S \<and> (\<forall>x \<in> 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.\<close>
definition\<^marker>\<open>tag important\<close> C1_differentiable_on :: "(real \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> real set \<Rightarrow> bool"
- (infix "C1'_differentiable'_on" 50)
+ (infix \<open>C1'_differentiable'_on\<close> 50)
where
"f C1_differentiable_on S \<longleftrightarrow>
(\<exists>D. (\<forall>x \<in> S. (f has_vector_derivative (D x)) (at x)) \<and> continuous_on S D)"
@@ -3336,7 +3336,7 @@
definition\<^marker>\<open>tag important\<close> piecewise_C1_differentiable_on
- (infixr "piecewise'_C1'_differentiable'_on" 50)
+ (infixr \<open>piecewise'_C1'_differentiable'_on\<close> 50)
where "f piecewise_C1_differentiable_on i \<equiv>
continuous_on i f \<and>
(\<exists>S. finite S \<and> (f C1_differentiable_on (i - S)))"
--- 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 \<Rightarrow> 'a"
- fixes proj::"'a \<Rightarrow> 'b \<Rightarrow> 'c::heine_borel" (infixl "proj" 60)
+ fixes proj::"'a \<Rightarrow> 'b \<Rightarrow> 'c::heine_borel" (infixl \<open>proj\<close> 60)
fixes unproj:: "('b \<Rightarrow> 'c) \<Rightarrow> 'a"
assumes finite_basis: "finite basis"
assumes bounded_proj: "\<And>k. k \<in> basis \<Longrightarrow> bounded ((\<lambda>x. x proj k) ` range f)"
--- 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 \<open>Filters\<close>
-definition indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter" (infixr "indirection" 70)
+definition indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter" (infixr \<open>indirection\<close> 70)
where "a indirection v = at a within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
--- 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 \<open>Limit Points\<close>
-definition\<^marker>\<open>tag important\<close> (in topological_space) islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60)
+definition\<^marker>\<open>tag important\<close> (in topological_space) islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr \<open>islimpt\<close> 60)
where "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))"
lemma islimptI:
@@ -2221,7 +2221,7 @@
qed
definition\<^marker>\<open>tag important\<close> homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
- (infixr "homeomorphic" 60)
+ (infixr \<open>homeomorphic\<close> 60)
where "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"
lemma homeomorphic_empty [iff]:
--- 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 \<Rightarrow> 'b::{banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
- (infixr "absolutely'_integrable'_on" 46)
+ (infixr \<open>absolutely'_integrable'_on\<close> 46)
where "f absolutely_integrable_on s \<equiv> set_integrable lebesgue s f"
--- 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\<open>HOL Light measurability\<close>
definition measurable_on :: "('a::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool"
- (infixr "measurable'_on" 46)
+ (infixr \<open>measurable'_on\<close> 46)
where "f measurable_on S \<equiv>
\<exists>N g. negligible N \<and>
(\<forall>n. continuous_on UNIV (g n)) \<and>
--- 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:
"(\<forall>u\<in>Basis. inner x u = 0) \<longleftrightarrow> (x = 0)"
-syntax "_type_dimension" :: "type \<Rightarrow> nat" ("(1DIM/(1'(_')))")
+syntax "_type_dimension" :: "type \<Rightarrow> nat" (\<open>(1DIM/(1'(_')))\<close>)
syntax_consts "_type_dimension" \<rightleftharpoons> card
translations "DIM('a)" \<rightharpoonup> "CONST card (CONST Basis :: 'a set)"
typed_print_translation \<open>
--- 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 @@
\<close>
definition\<^marker>\<open>tag important\<close>
has_fps_expansion :: "('a :: {banach,real_normed_div_algebra} \<Rightarrow> 'a) \<Rightarrow> 'a fps \<Rightarrow> bool"
- (infixl "has'_fps'_expansion" 60)
+ (infixl \<open>has'_fps'_expansion\<close> 60)
where "(f has_fps_expansion F) \<longleftrightarrow>
fps_conv_radius F > 0 \<and> eventually (\<lambda>z. eval_fps F z = f z) (nhds 0)"
--- 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 "\<chi>" 10)
+ vec_nth (infixl \<open>$\<close> 90) and
+ vec_lambda (binder \<open>\<chi>\<close> 10)
end
bundle no_vec_syntax begin
no_notation
- vec_nth (infixl "$" 90) and
- vec_lambda (binder "\<chi>" 10)
+ vec_nth (infixl \<open>$\<close> 90) and
+ vec_lambda (binder \<open>\<chi>\<close> 10)
end
unbundle vec_syntax
@@ -39,7 +39,7 @@
\<^item> \<open>'a^'b\<close> becomes \<open>('a, 'b::finite) vec\<close>
\<^item> \<open>'a^'b::_\<close> becomes \<open>('a, 'b) vec\<close> without extra sort-constraint
\<close>
-syntax "_vec_type" :: "type \<Rightarrow> type \<Rightarrow> type" (infixl "^" 15)
+syntax "_vec_type" :: "type \<Rightarrow> type \<Rightarrow> type" (infixl \<open>^\<close> 15)
syntax_types "_vec_type" \<rightleftharpoons> vec
parse_translation \<open>
let
@@ -283,7 +283,7 @@
text\<open>Also the scalar-vector multiplication.\<close>
-definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ 'n" (infixl "*s" 70)
+definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ 'n" (infixl \<open>*s\<close> 70)
where "c *s x = (\<chi> i. c * (x$i))"
text \<open>scalar product\<close>
@@ -986,15 +986,15 @@
by (simp add: map_matrix_def)
definition\<^marker>\<open>tag important\<close> matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
- (infixl "**" 70)
+ (infixl \<open>**\<close> 70)
where "m ** m' == (\<chi> i j. sum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
definition\<^marker>\<open>tag important\<close> matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
- (infixl "*v" 70)
+ (infixl \<open>*v\<close> 70)
where "m *v x \<equiv> (\<chi> i. sum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
definition\<^marker>\<open>tag important\<close> vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
- (infixl "v*" 70)
+ (infixl \<open>v*\<close> 70)
where "v v* m == (\<chi> j. sum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
definition\<^marker>\<open>tag unimportant\<close> "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
--- 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 \<equiv> PiM I M"
syntax
- "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>M _\<in>_./ _)" 10)
+ "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" (\<open>(3\<Pi>\<^sub>M _\<in>_./ _)\<close> 10)
syntax_consts
"_PiM" == PiM
translations
--- 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.\<close>
definition has_integral :: "('n::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> 'n set \<Rightarrow> bool"
- (infixr "has'_integral" 46)
+ (infixr \<open>has'_integral\<close> 46)
where "(f has_integral I) s \<longleftrightarrow>
(if \<exists>a b. s = cbox a b
then ((\<lambda>p. \<Sum>(x,k)\<in>p. content k *\<^sub>R f x) \<longlongrightarrow> I) (division_filter s)
@@ -321,7 +321,7 @@
(\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) (cbox a b) \<and> 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 \<open>integrable'_on\<close> 46)
where "f integrable_on i \<longleftrightarrow> (\<exists>y. (f has_integral y) i)"
definition "integral i f = (SOME y. (f has_integral y) i \<or> \<not> f integrable_on i \<and> y=0)"
--- 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\<open>Homotopy equivalence of topological spaces.\<close>
definition\<^marker>\<open>tag important\<close> homotopy_equivalent_space
- (infix "homotopy'_equivalent'_space" 50)
+ (infix \<open>homotopy'_equivalent'_space\<close> 50)
where "X homotopy_equivalent_space Y \<equiv>
(\<exists>f g. continuous_map X Y f \<and>
continuous_map Y X g \<and>
@@ -3781,7 +3781,7 @@
abbreviation\<^marker>\<open>tag important\<close> homotopy_eqv :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"
- (infix "homotopy'_eqv" 50)
+ (infix \<open>homotopy'_eqv\<close> 50)
where "S homotopy_eqv T \<equiv> top_of_set S homotopy_equivalent_space top_of_set T"
lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \<Longrightarrow> S homotopy_eqv T"
--- 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\<open>The definition here only really makes sense for an elementary set.
We just use compact intervals in applications below.\<close>
-definition\<^marker>\<open>tag important\<close> equiintegrable_on (infixr "equiintegrable'_on" 46)
+definition\<^marker>\<open>tag important\<close> equiintegrable_on (infixr \<open>equiintegrable'_on\<close> 46)
where "F equiintegrable_on I \<equiv>
(\<forall>f \<in> F. f integrable_on I) \<and>
(\<forall>e > 0. \<exists>\<gamma>. gauge \<gamma> \<and>
--- 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\<open>The nonzero and zero cases, as in \emph{Complex Analysis} by Joseph Bak and Donald J.Newman, page 241\<close>
text\<^marker>\<open>tag important\<close> \<open>%whitespace\<close>
definition\<^marker>\<open>tag important\<close>
- has_prod :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "has'_prod" 80)
+ has_prod :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a \<Rightarrow> bool" (infixr \<open>has'_prod\<close> 80)
where "f has_prod p \<equiv> raw_has_prod f 0 p \<or> (\<exists>i q. p = 0 \<and> f i = 0 \<and> raw_has_prod f (Suc i) q)"
definition\<^marker>\<open>tag important\<close> convergent_prod :: "(nat \<Rightarrow> 'a :: {t2_space,comm_semiring_1}) \<Rightarrow> bool" where
"convergent_prod f \<equiv> \<exists>M p. raw_has_prod f M p"
definition\<^marker>\<open>tag important\<close> prodinf :: "(nat \<Rightarrow> 'a::{t2_space, comm_semiring_1}) \<Rightarrow> 'a"
- (binder "\<Prod>" 10)
+ (binder \<open>\<Prod>\<close> 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
--- 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 \<open>Conflicting notation from \<^theory>\<open>HOL-Analysis.Infinite_Sum\<close>\<close>
-no_notation Infinite_Sum.abs_summable_on (infixr "abs'_summable'_on" 46)
+no_notation Infinite_Sum.abs_summable_on (infixr \<open>abs'_summable'_on\<close> 46)
(* TODO Move *)
lemma sets_eq_countable:
@@ -84,7 +84,7 @@
definition\<^marker>\<open>tag important\<close> abs_summable_on ::
"('a \<Rightarrow> 'b :: {banach, second_countable_topology}) \<Rightarrow> 'a set \<Rightarrow> bool"
- (infix "abs'_summable'_on" 50)
+ (infix \<open>abs'_summable'_on\<close> 50)
where
"f abs_summable_on A \<longleftrightarrow> integrable (count_space A) f"
@@ -96,10 +96,10 @@
syntax (ASCII)
"_infsetsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
- ("(3INFSETSUM _:_./ _)" [0, 51, 10] 10)
+ (\<open>(3INFSETSUM _:_./ _)\<close> [0, 51, 10] 10)
syntax
"_infsetsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
- ("(2\<Sum>\<^sub>a_\<in>_./ _)" [0, 51, 10] 10)
+ (\<open>(2\<Sum>\<^sub>a_\<in>_./ _)\<close> [0, 51, 10] 10)
syntax_consts
"_infsetsum" \<rightleftharpoons> infsetsum
translations \<comment> \<open>Beware of argument permutation!\<close>
@@ -107,10 +107,10 @@
syntax (ASCII)
"_uinfsetsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
- ("(3INFSETSUM _:_./ _)" [0, 51, 10] 10)
+ (\<open>(3INFSETSUM _:_./ _)\<close> [0, 51, 10] 10)
syntax
"_uinfsetsum" :: "pttrn \<Rightarrow> 'b \<Rightarrow> 'b::{banach, second_countable_topology}"
- ("(2\<Sum>\<^sub>a_./ _)" [0, 10] 10)
+ (\<open>(2\<Sum>\<^sub>a_./ _)\<close> [0, 10] 10)
syntax_consts
"_uinfsetsum" \<rightleftharpoons> infsetsum
translations \<comment> \<open>Beware of argument permutation!\<close>
@@ -118,10 +118,10 @@
syntax (ASCII)
"_qinfsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a::{banach, second_countable_topology}"
- ("(3INFSETSUM _ |/ _./ _)" [0, 0, 10] 10)
+ (\<open>(3INFSETSUM _ |/ _./ _)\<close> [0, 0, 10] 10)
syntax
"_qinfsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a::{banach, second_countable_topology}"
- ("(2\<Sum>\<^sub>a_ | (_)./ _)" [0, 0, 10] 10)
+ (\<open>(2\<Sum>\<^sub>a_ | (_)./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_qinfsetsum" \<rightleftharpoons> infsetsum
translations
--- 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 :: \<open>('a \<Rightarrow> 'b :: {comm_monoid_add, topological_space}) \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool\<close>
where has_sum_def: \<open>HAS_SUM f A x \<equiv> (sum f \<longlongrightarrow> x) (finite_subsets_at_top A)\<close>
-abbreviation has_sum (infixr "has'_sum" 46) where
+abbreviation has_sum (infixr \<open>has'_sum\<close> 46) where
"(f has_sum S) A \<equiv> HAS_SUM f A S"
-definition summable_on :: "('a \<Rightarrow> 'b::{comm_monoid_add, topological_space}) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "summable'_on" 46) where
+definition summable_on :: "('a \<Rightarrow> 'b::{comm_monoid_add, topological_space}) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr \<open>summable'_on\<close> 46) where
"f summable_on A \<equiv> (\<exists>x. (f has_sum x) A)"
definition infsum :: "('a \<Rightarrow> 'b::{comm_monoid_add,t2_space}) \<Rightarrow> 'a set \<Rightarrow> '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 \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "abs'_summable'_on" 46) where
+abbreviation abs_summable_on :: "('a \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a set \<Rightarrow> bool" (infixr \<open>abs'_summable'_on\<close> 46) where
"f abs_summable_on A \<equiv> (\<lambda>x. norm (f x)) summable_on A"
syntax (ASCII)
- "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add" ("(3INFSUM (_/:_)./ _)" [0, 51, 10] 10)
+ "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add" (\<open>(3INFSUM (_/:_)./ _)\<close> [0, 51, 10] 10)
syntax
- "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add" ("(2\<Sum>\<^sub>\<infinity>(_/\<in>_)./ _)" [0, 51, 10] 10)
+ "_infsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::topological_comm_monoid_add" (\<open>(2\<Sum>\<^sub>\<infinity>(_/\<in>_)./ _)\<close> [0, 51, 10] 10)
syntax_consts
"_infsum" \<rightleftharpoons> infsum
translations \<comment> \<open>Beware of argument permutation!\<close>
"\<Sum>\<^sub>\<infinity>i\<in>A. b" \<rightleftharpoons> "CONST infsum (\<lambda>i. b) A"
syntax (ASCII)
- "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a" ("(3INFSUM _./ _)" [0, 10] 10)
+ "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(3INFSUM _./ _)\<close> [0, 10] 10)
syntax
- "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a" ("(2\<Sum>\<^sub>\<infinity>_./ _)" [0, 10] 10)
+ "_univinfsum" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(2\<Sum>\<^sub>\<infinity>_./ _)\<close> [0, 10] 10)
syntax_consts
"_univinfsum" \<rightleftharpoons> infsum
translations
"\<Sum>\<^sub>\<infinity>x. t" \<rightleftharpoons> "CONST infsum (\<lambda>x. t) (CONST UNIV)"
syntax (ASCII)
- "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3INFSUM _ |/ _./ _)" [0, 0, 10] 10)
+ "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(3INFSUM _ |/ _./ _)\<close> [0, 0, 10] 10)
syntax
- "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(2\<Sum>\<^sub>\<infinity>_ | (_)./ _)" [0, 0, 10] 10)
+ "_qinfsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(2\<Sum>\<^sub>\<infinity>_ | (_)./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_qinfsum" \<rightleftharpoons> infsum
translations
--- 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>\<open>tag important\<close>
gderiv ::
"['a::real_inner \<Rightarrow> real, 'a, 'a] \<Rightarrow> bool"
- ("(GDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
+ (\<open>(GDERIV (_)/ (_)/ :> (_))\<close> [1000, 1000, 60] 60)
where
"GDERIV f x :> D \<longleftrightarrow> FDERIV f x :> (\<lambda>h. inner h D)"
@@ -463,11 +463,11 @@
lemmas has_derivative_norm = GDERIV_norm [unfolded gderiv_def]
bundle inner_syntax begin
-notation inner (infix "\<bullet>" 70)
+notation inner (infix \<open>\<bullet>\<close> 70)
end
bundle no_inner_syntax begin
-no_notation inner (infix "\<bullet>" 70)
+no_notation inner (infix \<open>\<bullet>\<close> 70)
end
end
--- 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 \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real measure \<Rightarrow> real \<Rightarrow> real"
- ("(5LINT _=_.._|_. _)" [0,60,60,61,100] 60)
+ (\<open>(5LINT _=_.._|_. _)\<close> [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 \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real"
- ("(4LBINT _=_.._. _)" [0,60,60,61] 60)
+ (\<open>(4LBINT _=_.._. _)\<close> [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 \<Rightarrow> real \<Rightarrow> complex"
- ("(2CLBINT _. _)" [0,60] 60)
+ (\<open>(2CLBINT _. _)\<close> [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 (\<lambda>x. f)"
syntax "_complex_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> complex"
- ("(3CLBINT _:_. _)" [0,60,61] 60)
+ (\<open>(3CLBINT _:_. _)\<close> [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 \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> real \<Rightarrow> complex"
- ("(4CLBINT _=_.._. _)" [0,60,60,61] 60)
+ (\<open>(4CLBINT _=_.._. _)\<close> [0,60,60,61] 60)
syntax_consts
"_ascii_complex_interval_lebesgue_borel_integral" == complex_interval_lebesgue_integral
--- 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 \<open>Isolate and discrete\<close>
-definition (in topological_space) isolated_in:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "isolated'_in" 60)
+definition (in topological_space) isolated_in:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr \<open>isolated'_in\<close> 60)
where "x isolated_in S \<longleftrightarrow> (x\<in>S \<and> (\<exists>T. open T \<and> T \<inter> S = {x}))"
definition (in topological_space) discrete:: "'a set \<Rightarrow> bool"
--- 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 \<longleftrightarrow> (0 \<le> C \<and> (\<forall>x \<in> U. \<forall>y\<in>U. dist (f x) (f y) \<le> C * dist x y))"
bundle lipschitz_syntax begin
-notation\<^marker>\<open>tag important\<close> lipschitz_on ("_-lipschitz'_on" [1000])
+notation\<^marker>\<open>tag important\<close> lipschitz_on (\<open>_-lipschitz'_on\<close> [1000])
end
bundle no_lipschitz_syntax begin
-no_notation lipschitz_on ("_-lipschitz'_on" [1000])
+no_notation lipschitz_on (\<open>_-lipschitz'_on\<close> [1000])
end
unbundle lipschitz_syntax
--- 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\<open>Basic definition of the small inductive dimension relation. Works in any topological space.\<close>
-inductive dimension_le :: "['a topology, int] \<Rightarrow> bool" (infix "dim'_le" 50)
+inductive dimension_le :: "['a topology, int] \<Rightarrow> bool" (infix \<open>dim'_le\<close> 50)
where "\<lbrakk>-1 \<le> n;
\<And>V a. \<lbrakk>openin X V; a \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. a \<in> U \<and> U \<subseteq> V \<and> openin X U \<and> (subtopology X (X frontier_of U)) dim_le (n-1)\<rbrakk>
\<Longrightarrow> X dim_le (n::int)"
--- 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 \<equiv> eventually P (ae_filter M)"
syntax
- "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
+ "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" (\<open>AE _ in _. _\<close> [0,0,10] 10)
syntax_consts
"_almost_everywhere" \<rightleftharpoons> almost_everywhere
@@ -1017,7 +1017,7 @@
syntax
"_set_almost_everywhere" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool"
- ("AE _\<in>_ in _./ _" [0,0,0,10] 10)
+ (\<open>AE _\<in>_ in _./ _\<close> [0,0,0,10] 10)
syntax_consts
"_set_almost_everywhere" \<rightleftharpoons> set_almost_everywhere
--- 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>\<open>tag important\<close> simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>S") where
+definition\<^marker>\<open>tag important\<close> simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" (\<open>integral\<^sup>S\<close>) where
"integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
syntax
- "_simple_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
+ "_simple_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" (\<open>\<integral>\<^sup>S _. _ \<partial>_\<close> [60,61] 110)
syntax_consts
"_simple_integral" == simple_integral
@@ -816,11 +816,11 @@
subsection \<open>Integral on nonnegative functions\<close>
-definition\<^marker>\<open>tag important\<close> nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>N") where
+definition\<^marker>\<open>tag important\<close> nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" (\<open>integral\<^sup>N\<close>) where
"integral\<^sup>N M f = (SUP g \<in> {g. simple_function M g \<and> g \<le> f}. integral\<^sup>S M g)"
syntax
- "_nn_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>+((2 _./ _)/ \<partial>_)" [60,61] 110)
+ "_nn_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" (\<open>\<integral>\<^sup>+((2 _./ _)/ \<partial>_)\<close> [60,61] 110)
syntax_consts
"_nn_integral" == nn_integral
--- 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 "<e" 50)
+ eucl_less (infix \<open><e\<close> 50)
lemma One_nonneg: "0 \<le> (\<Sum>Basis::'a::ordered_euclidean_space)"
by (auto intro: sum_nonneg)
--- 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 \<equiv> (\<lambda>x. g(1 - x))"
definition\<^marker>\<open>tag important\<close> joinpaths :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a"
- (infixr "+++" 75)
+ (infixr \<open>+++\<close> 75)
where "g1 +++ g2 \<equiv> (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
definition\<^marker>\<open>tag important\<close> loop_free :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
--- 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 \<open>Faces of a (usually convex) set\<close>
-definition\<^marker>\<open>tag important\<close> face_of :: "['a::real_vector set, 'a set] \<Rightarrow> bool" (infixr "(face'_of)" 50)
+definition\<^marker>\<open>tag important\<close> face_of :: "['a::real_vector set, 'a set] \<Rightarrow> bool" (infixr \<open>(face'_of)\<close> 50)
where
"T face_of S \<longleftrightarrow>
T \<subseteq> S \<and> convex T \<and>
@@ -654,7 +654,7 @@
text\<open>That is, faces that are intersection with supporting hyperplane\<close>
definition\<^marker>\<open>tag important\<close> exposed_face_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
- (infixr "(exposed'_face'_of)" 50)
+ (infixr \<open>(exposed'_face'_of)\<close> 50)
where "T exposed_face_of S \<longleftrightarrow>
T face_of S \<and> (\<exists>a b. S \<subseteq> {x. a \<bullet> x \<le> b} \<and> T = S \<inter> {x. a \<bullet> x = b})"
@@ -873,7 +873,7 @@
subsection\<open>Extreme points of a set: its singleton faces\<close>
definition\<^marker>\<open>tag important\<close> extreme_point_of :: "['a::real_vector, 'a set] \<Rightarrow> bool"
- (infixr "(extreme'_point'_of)" 50)
+ (infixr \<open>(extreme'_point'_of)\<close> 50)
where "x extreme_point_of S \<longleftrightarrow>
x \<in> S \<and> (\<forall>a \<in> S. \<forall>b \<in> S. x \<notin> open_segment a b)"
@@ -979,7 +979,7 @@
subsection\<open>Facets\<close>
definition\<^marker>\<open>tag important\<close> facet_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool"
- (infixr "(facet'_of)" 50)
+ (infixr \<open>(facet'_of)\<close> 50)
where "F facet_of S \<longleftrightarrow> F face_of S \<and> F \<noteq> {} \<and> aff_dim F = aff_dim S - 1"
lemma facet_of_empty [simp]: "\<not> S facet_of {}"
@@ -1024,7 +1024,7 @@
subsection \<open>Edges: faces of affine dimension 1\<close> (*FIXME too small subsection, rearrange? *)
-definition\<^marker>\<open>tag important\<close> edge_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool" (infixr "(edge'_of)" 50)
+definition\<^marker>\<open>tag important\<close> edge_of :: "['a::euclidean_space set, 'a set] \<Rightarrow> bool" (infixr \<open>(edge'_of)\<close> 50)
where "e edge_of S \<longleftrightarrow> e face_of S \<and> aff_dim e = 1"
lemma edge_of_imp_subset:
@@ -2984,7 +2984,7 @@
text\<open>The notion of n-simplex for integer \<^term>\<open>n \<ge> -1\<close>\<close>
-definition\<^marker>\<open>tag important\<close> simplex :: "int \<Rightarrow> 'a::euclidean_space set \<Rightarrow> bool" (infix "simplex" 50)
+definition\<^marker>\<open>tag important\<close> simplex :: "int \<Rightarrow> 'a::euclidean_space set \<Rightarrow> bool" (infix \<open>simplex\<close> 50)
where "n simplex S \<equiv> \<exists>C. \<not> affine_dependent C \<and> int(card C) = n + 1 \<and> S = convex hull C"
lemma simplex:
--- 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 \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
- ("(4LINT (_):(_)/|(_)./ _)" [0,0,0,10] 10)
+ (\<open>(4LINT (_):(_)/|(_)./ _)\<close> [0,0,0,10] 10)
syntax_consts
"_ascii_set_lebesgue_integral" == set_lebesgue_integral
@@ -43,11 +43,11 @@
syntax
"_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real"
- ("(2LBINT _./ _)" [0,10] 10)
+ (\<open>(2LBINT _./ _)\<close> [0,10] 10)
syntax
"_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> real"
- ("(3LBINT _:_./ _)" [0,0,10] 10)
+ (\<open>(3LBINT _:_./ _)\<close> [0,0,10] 10)
section \<open>Basic properties\<close>
@@ -572,12 +572,12 @@
abbreviation complex_integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> bool" where
"complex_integrable M f \<equiv> integrable M f"
-abbreviation complex_lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" ("integral\<^sup>C") where
+abbreviation complex_lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" (\<open>integral\<^sup>C\<close>) where
"integral\<^sup>C M f == integral\<^sup>L M f"
syntax
"_complex_lebesgue_integral" :: "pttrn \<Rightarrow> complex \<Rightarrow> 'a measure \<Rightarrow> complex"
- ("\<integral>\<^sup>C _. _ \<partial>_" [0,0] 110)
+ (\<open>\<integral>\<^sup>C _. _ \<partial>_\<close> [0,0] 110)
syntax_consts
"_complex_lebesgue_integral" == complex_lebesgue_integral
@@ -587,7 +587,7 @@
syntax
"_ascii_complex_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
- ("(3CLINT _|_. _)" [0,0,10] 10)
+ (\<open>(3CLINT _|_. _)\<close> [0,0,10] 10)
syntax_consts
"_ascii_complex_lebesgue_integral" == complex_lebesgue_integral
@@ -624,7 +624,7 @@
syntax
"_ascii_complex_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
- ("(4CLINT _:_|_. _)" [0,0,0,10] 10)
+ (\<open>(4CLINT _:_|_. _)\<close> [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"
-("(\<integral>\<^sup>+((_)\<in>(_)./ _)/\<partial>_)" [0,0,0,110] 10)
+(\<open>(\<integral>\<^sup>+((_)\<in>(_)./ _)/\<partial>_)\<close> [0,0,0,110] 10)
"_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal"
-("(\<integral>((_)\<in>(_)./ _)/\<partial>_)" [0,0,0,110] 10)
+(\<open>(\<integral>((_)\<in>(_)./ _)/\<partial>_)\<close> [0,0,0,110] 10)
syntax_consts
"_set_nn_integral" == set_nn_integral and
--- 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 \<open>Measurable functions\<close>
definition\<^marker>\<open>tag important\<close> measurable :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) set"
- (infixr "\<rightarrow>\<^sub>M" 60) where
+ (infixr \<open>\<rightarrow>\<^sub>M\<close> 60) where
"measurable A B = {f \<in> space A \<rightarrow> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
lemma measurableI:
--- 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 \<open>A set of points sparse in another set\<close>
definition sparse_in:: "'a :: topological_space set \<Rightarrow> 'a set \<Rightarrow> bool"
- (infixl "(sparse'_in)" 50)
+ (infixl \<open>(sparse'_in)\<close> 50)
where
"pts sparse_in A = (\<forall>x\<in>A. \<exists>B. x\<in>B \<and> open B \<and> (\<forall>y\<in>B. \<not> y islimpt pts))"
@@ -165,14 +165,14 @@
"cosparse A = Abs_filter (\<lambda>P. {x. \<not>P x} sparse_in A)"
syntax
- "_eventually_cosparse" :: "pttrn => 'a set => bool => bool" ("(3\<forall>\<^sub>\<approx>_\<in>_./ _)" [0, 0, 10] 10)
+ "_eventually_cosparse" :: "pttrn => 'a set => bool => bool" (\<open>(3\<forall>\<^sub>\<approx>_\<in>_./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_eventually_cosparse" == eventually
translations
"\<forall>\<^sub>\<approx>x\<in>A. P" == "CONST eventually (\<lambda>x. P) (CONST cosparse A)"
syntax
- "_qeventually_cosparse" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<forall>\<^sub>\<approx>_ | (_)./ _)" [0, 0, 10] 10)
+ "_qeventually_cosparse" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(3\<forall>\<^sub>\<approx>_ | (_)./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_qeventually_cosparse" == eventually
translations
--- 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\<open>Orthogonal complement\<close>
-definition\<^marker>\<open>tag important\<close> orthogonal_comp ("_\<^sup>\<bottom>" [80] 80)
+definition\<^marker>\<open>tag important\<close> orthogonal_comp (\<open>_\<^sup>\<bottom>\<close> [80] 80)
where "orthogonal_comp W \<equiv> {x. \<forall>y \<in> W. orthogonal y x}"
proposition subspace_orthogonal_comp: "subspace (W\<^sup>\<bottom>)"
--- 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 \<open>Divisions\<close>
-definition\<^marker>\<open>tag important\<close> division_of (infixl "division'_of" 40)
+definition\<^marker>\<open>tag important\<close> division_of (infixl \<open>division'_of\<close> 40)
where
"s division_of i \<longleftrightarrow>
finite s \<and>
@@ -820,7 +820,7 @@
subsection \<open>Tagged (partial) divisions\<close>
-definition\<^marker>\<open>tag important\<close> tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
+definition\<^marker>\<open>tag important\<close> tagged_partial_division_of (infixr \<open>tagged'_partial'_division'_of\<close> 40)
where "s tagged_partial_division_of i \<longleftrightarrow>
finite s \<and>
(\<forall>x K. (x, K) \<in> s \<longrightarrow> x \<in> K \<and> K \<subseteq> i \<and> (\<exists>a b. K = cbox a b)) \<and>
@@ -837,7 +837,7 @@
(x2, K2) \<in> s \<Longrightarrow> (x1, K1) \<noteq> (x2, K2) \<Longrightarrow> interior K1 \<inter> interior K2 = {}"
using assms unfolding tagged_partial_division_of_def by blast+
-definition\<^marker>\<open>tag important\<close> tagged_division_of (infixr "tagged'_division'_of" 40)
+definition\<^marker>\<open>tag important\<close> tagged_division_of (infixr \<open>tagged'_division'_of\<close> 40)
where "s tagged_division_of i \<longleftrightarrow> s tagged_partial_division_of i \<and> (\<Union>{K. \<exists>x. (x,K) \<in> s} = i)"
lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s"
@@ -1593,7 +1593,7 @@
subsection \<open>Fine-ness of a partition w.r.t. a gauge\<close>
-definition\<^marker>\<open>tag important\<close> fine (infixr "fine" 46)
+definition\<^marker>\<open>tag important\<close> fine (infixr \<open>fine\<close> 46)
where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)"
lemma fineI:
--- 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>\<open>tag unimportant\<close> Zero_neq_One[iff]: "0 \<noteq> One"
by (metis One_non_0)
-definition\<^marker>\<open>tag important\<close> (in euclidean_space) eucl_less (infix "<e" 50) where
+definition\<^marker>\<open>tag important\<close> (in euclidean_space) eucl_less (infix \<open><e\<close> 50) where
"eucl_less a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i)"
definition\<^marker>\<open>tag important\<close> box_eucl_less: "box a b = {x. a <e x \<and> x <e b}"
@@ -2485,6 +2485,6 @@
qed auto
no_notation
- eucl_less (infix "<e" 50)
+ eucl_less (infix \<open><e\<close> 50)
end
--- 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] \<Rightarrow> bool"
- ("_ Issues _ with _ on _" [50, 0, 0, 50] 50) where
+ (\<open>_ Issues _ with _ on _\<close> [50, 0, 0, 50] 50) where
"(A Issues B with X on evs) =
(\<exists>Y. Says A B Y \<in> set evs \<and> X \<in> parts {Y} \<and>
X \<notin> parts (spies (takeWhile (\<lambda>z. z \<noteq> 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] \<Rightarrow> event list" ("before _ on _" [0, 50] 50)
+ before :: "[event, event list] \<Rightarrow> event list" (\<open>before _ on _\<close> [0, 50] 50)
where "(before ev on evs) = takeWhile (\<lambda>z. z \<noteq> ev) (rev evs)"
definition
(* States than an event really appears only once on a trace *)
- Unique :: "[event, event list] \<Rightarrow> bool" ("Unique _ on _" [0, 50] 50)
+ Unique :: "[event, event list] \<Rightarrow> bool" (\<open>Unique _ on _\<close> [0, 50] 50)
where "(Unique ev on evs) = (ev \<notin> set (tl (dropWhile (\<lambda>z. z \<noteq> ev) evs)))"
@@ -95,7 +95,7 @@
"expiredA T evs == authlife + T < CT evs"
abbreviation
- valid :: "[nat, nat] \<Rightarrow> bool" ("valid _ wrt _" [0, 50] 50) where
+ valid :: "[nat, nat] \<Rightarrow> bool" (\<open>valid _ wrt _\<close> [0, 50] 50) where
"valid T1 wrt T2 == T1 \<le> replylife + T2"
(*---------------------------------------------------------------------*)
--- 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] \<Rightarrow> bool" ("Unique _ on _" [0, 50] 50)
+ Unique :: "[event, event list] \<Rightarrow> bool" (\<open>Unique _ on _\<close> [0, 50] 50)
where "(Unique ev on evs) = (ev \<notin> set (tl (dropWhile (\<lambda>z. z \<noteq> ev) evs)))"
@@ -81,7 +81,7 @@
"expiredA T evs == authlife + T < CT evs"
abbreviation
- valid :: "[nat, nat] \<Rightarrow> bool" ("valid _ wrt _" [0, 50] 50) where
+ valid :: "[nat, nat] \<Rightarrow> bool" (\<open>valid _ wrt _\<close> [0, 50] 50) where
"valid T1 wrt T2 == T1 \<le> replylife + T2"
(*---------------------------------------------------------------------*)
--- 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] \<Rightarrow> bool"
- ("_ Issues _ with _ on _") where
+ (\<open>_ Issues _ with _ on _\<close>) where
"A Issues B with X on evs =
(\<exists>Y. Says A B Y \<in> set evs \<and> X \<in> parts {Y} \<and>
X \<notin> parts (spies (takeWhile (\<lambda>z. z \<noteq> Says A B Y) (rev evs))))"
@@ -86,7 +86,7 @@
"expiredA T evs == authlife + T < CT evs"
abbreviation
- valid :: "[nat, nat] \<Rightarrow> bool" ("valid _ wrt _") where
+ valid :: "[nat, nat] \<Rightarrow> bool" (\<open>valid _ wrt _\<close>) where
"valid T1 wrt T2 == T1 \<le> replylife + T2"
(*---------------------------------------------------------------------*)
--- 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] \<Rightarrow> bool"
- ("_ Issues _ with _ on _") where
+ (\<open>_ Issues _ with _ on _\<close>) where
"A Issues B with X on evs =
(\<exists>Y. Says A B Y \<in> set evs \<and> X \<in> parts {Y} \<and>
X \<notin> parts (spies (takeWhile (\<lambda>z. z \<noteq> 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] \<Rightarrow> event list" ("before _ on _")
+ before :: "[event, event list] \<Rightarrow> event list" (\<open>before _ on _\<close>)
where "before ev on evs = takeWhile (\<lambda>z. z \<noteq> ev) (rev evs)"
definition
(* States than an event really appears only once on a trace *)
- Unique :: "[event, event list] \<Rightarrow> bool" ("Unique _ on _")
+ Unique :: "[event, event list] \<Rightarrow> bool" (\<open>Unique _ on _\<close>)
where "Unique ev on evs = (ev \<notin> set (tl (dropWhile (\<lambda>z. z \<noteq> ev) evs)))"
--- 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] \<Rightarrow> event list" ("before _ on _")
+ before :: "[event, event list] \<Rightarrow> event list" (\<open>before _ on _\<close>)
where "before ev on evs = takeWhile (\<lambda>z. z \<noteq> ev) (rev evs)"
definition
(* States than an event really appears only once on a trace *)
- Unique :: "[event, event list] \<Rightarrow> bool" ("Unique _ on _")
+ Unique :: "[event, event list] \<Rightarrow> bool" (\<open>Unique _ on _\<close>)
where "Unique ev on evs = (ev \<notin> set (tl (dropWhile (\<lambda>z. z \<noteq> ev) evs)))"
--- 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\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
nonterminal mtuple_args
syntax
- "" :: "'a \<Rightarrow> mtuple_args" ("_")
- "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" ("_,/ _")
- "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)")
+ "" :: "'a \<Rightarrow> mtuple_args" (\<open>_\<close>)
+ "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" (\<open>_,/ _\<close>)
+ "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" (\<open>(2\<lbrace>_,/ _\<rbrace>)\<close>)
syntax_consts
"_MTuple_args" "_MTuple" \<rightleftharpoons> MPair
translations
@@ -61,7 +61,7 @@
"\<lbrace>x, y\<rbrace>" \<rightleftharpoons> "CONST MPair x y"
-definition HPair :: "[msg,msg] \<Rightarrow> msg" ("(4Hash[_] /_)" [0, 1000]) where
+definition HPair :: "[msg,msg] \<Rightarrow> msg" (\<open>(4Hash[_] /_)\<close> [0, 1000]) where
\<comment> \<open>Message Y paired with a MAC computed with the help of X\<close>
"Hash[X] Y == \<lbrace>Hash\<lbrace>X,Y\<rbrace>, Y\<rbrace>"
--- 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] \<Rightarrow> bool"
- ("_ Issues _ with _ on _") where
+ (\<open>_ Issues _ with _ on _\<close>) where
"A Issues B with X on evs =
(\<exists>Y. Says A B Y \<in> set evs \<and> X \<in> parts {Y} \<and>
X \<notin> parts (spies (takeWhile (\<lambda>z. z \<noteq> Says A B Y) (rev evs))))"
--- 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 \<in> sr \<Longrightarrow> secureM"
-definition Unique :: "[event, event list] => bool" ("Unique _ on _") where
+definition Unique :: "[event, event list] => bool" (\<open>Unique _ on _\<close>) where
"Unique ev on evs ==
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
--- 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 \<in> srb \<Longrightarrow> secureM"
-definition Unique :: "[event, event list] => bool" ("Unique _ on _") where
+definition Unique :: "[event, event list] => bool" (\<open>Unique _ on _\<close>) where
"Unique ev on evs ==
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
--- 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 \<noteq> pin Q" and
crdK_disj_pin [iff]: "crdK C \<noteq> pin P"
-definition legalUse :: "card => bool" ("legalUse (_)") where
+definition legalUse :: "card => bool" (\<open>legalUse (_)\<close>) where
"legalUse C == C \<notin> stolen"
primrec illegalUse :: "card => bool" where
--- 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 \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> bool" ("_\<turnstile>init\<le>_" [51,51] 50)
+ init_le :: "prog \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> bool" (\<open>_\<turnstile>init\<le>_\<close> [51,51] 50)
where "G\<turnstile>init\<le>n = (\<lambda>s. card (nyinitcls G s) \<le> n)"
lemma init_le_def2 [simp]: "(G\<turnstile>init\<le>n) s = (card (nyinitcls G s)\<le>n)"
@@ -116,7 +116,7 @@
subsubsection "Most General Triples and Formulas"
definition
- remember_init_state :: "state assn" ("\<doteq>")
+ remember_init_state :: "state assn" (\<open>\<doteq>\<close>)
where "\<doteq> \<equiv> \<lambda>Y s Z. s = Z"
lemma remember_init_state_def2 [simp]: "\<doteq> Y = (=)"
@@ -125,11 +125,11 @@
done
definition
- MGF ::"[state assn, term, prog] \<Rightarrow> state triple" ("{_} _\<succ> {_\<rightarrow>}"[3,65,3]62)
+ MGF ::"[state assn, term, prog] \<Rightarrow> state triple" (\<open>{_} _\<succ> {_\<rightarrow>}\<close>[3,65,3]62)
where "{P} t\<succ> {G\<rightarrow>} = {P} t\<succ> {\<lambda>Y s' s. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (Y,s')}"
definition
- MGFn :: "[nat, term, prog] \<Rightarrow> state triple" ("{=:_} _\<succ> {_\<rightarrow>}"[3,65,3]62)
+ MGFn :: "[nat, term, prog] \<Rightarrow> state triple" (\<open>{=:_} _\<succ> {_\<rightarrow>}\<close>[3,65,3]62)
where "{=:n} t\<succ> {G\<rightarrow>} = {\<doteq> \<and>. G\<turnstile>init\<le>n} t\<succ> {G\<rightarrow>}"
(* unused *)
--- 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" (\<open>Val:_\<close> [951] 950)
+ "_Var" :: "[pttrn] => pttrn" (\<open>Var:_\<close> [951] 950)
+ "_Vals" :: "[pttrn] => pttrn" (\<open>Vals:_\<close> [951] 950)
translations
"\<lambda>Val:v . b" == "(\<lambda>v. b) \<circ> CONST the_In1"
@@ -63,7 +63,7 @@
(type) "'a assn" <= (type) "vals \<Rightarrow> state \<Rightarrow> 'a \<Rightarrow> bool"
definition
- assn_imp :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool" (infixr "\<Rightarrow>" 25)
+ assn_imp :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool" (infixr \<open>\<Rightarrow>\<close> 25)
where "(P \<Rightarrow> Q) = (\<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z)"
lemma assn_imp_def2 [iff]: "(P \<Rightarrow> Q) = (\<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z)"
@@ -77,7 +77,7 @@
subsection "peek-and"
definition
- peek_and :: "'a assn \<Rightarrow> (state \<Rightarrow> bool) \<Rightarrow> 'a assn" (infixl "\<and>." 13)
+ peek_and :: "'a assn \<Rightarrow> (state \<Rightarrow> bool) \<Rightarrow> 'a assn" (infixl \<open>\<and>.\<close> 13)
where "(P \<and>. p) = (\<lambda>Y s Z. P Y s Z \<and> p s)"
lemma peek_and_def2 [simp]: "peek_and P p Y s = (\<lambda>Z. (P Y s Z \<and> p s))"
@@ -117,7 +117,7 @@
subsection "assn-supd"
definition
- assn_supd :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl ";." 13)
+ assn_supd :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl \<open>;.\<close> 13)
where "(P ;. f) = (\<lambda>Y s' Z. \<exists>s. P Y s Z \<and> s' = f s)"
lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (\<exists>s. P Y s Z \<and> s' = f s)"
@@ -128,7 +128,7 @@
subsection "supd-assn"
definition
- supd_assn :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr ".;" 13)
+ supd_assn :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr \<open>.;\<close> 13)
where "(f .; P) = (\<lambda>Y s. P Y (f s))"
@@ -148,7 +148,7 @@
subsection "subst-res"
definition
- subst_res :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" ("_\<leftarrow>_" [60,61] 60)
+ subst_res :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" (\<open>_\<leftarrow>_\<close> [60,61] 60)
where "P\<leftarrow>w = (\<lambda>Y. P w)"
lemma subst_res_def2 [simp]: "(P\<leftarrow>w) Y = P w"
@@ -184,7 +184,7 @@
subsection "subst-Bool"
definition
- subst_Bool :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn" ("_\<leftarrow>=_" [60,61] 60)
+ subst_Bool :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn" (\<open>_\<leftarrow>=_\<close> [60,61] 60)
where "P\<leftarrow>=b = (\<lambda>Y s Z. \<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b))"
lemma subst_Bool_def2 [simp]:
@@ -204,7 +204,7 @@
where "peek_res Pf = (\<lambda>Y. Pf Y Y)"
syntax
- "_peek_res" :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" ("\<lambda>_:. _" [0,3] 3)
+ "_peek_res" :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (\<open>\<lambda>_:. _\<close> [0,3] 3)
syntax_consts
"_peek_res" == peek_res
translations
@@ -231,7 +231,7 @@
subsection "ign-res"
definition
- ign_res :: "'a assn \<Rightarrow> 'a assn" ("_\<down>" [1000] 1000)
+ ign_res :: "'a assn \<Rightarrow> 'a assn" (\<open>_\<down>\<close> [1000] 1000)
where "P\<down> = (\<lambda>Y s Z. \<exists>Y. P Y s Z)"
lemma ign_res_def2 [simp]: "P\<down> Y s Z = (\<exists>Y. P Y s Z)"
@@ -267,7 +267,7 @@
where "peek_st P = (\<lambda>Y s. P (store s) Y s)"
syntax
- "_peek_st" :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" ("\<lambda>_.. _" [0,3] 3)
+ "_peek_st" :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (\<open>\<lambda>_.. _\<close> [0,3] 3)
syntax_consts
"_peek_st" == peek_st
translations
@@ -310,7 +310,7 @@
subsection "ign-res-eq"
definition
- ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" ("_\<down>=_" [60,61] 60)
+ ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" (\<open>_\<down>=_\<close> [60,61] 60)
where "P\<down>=w \<equiv> (\<lambda>Y:. P\<down> \<and>. (\<lambda>s. Y=w))"
lemma ign_res_eq_def2 [simp]: "(P\<down>=w) Y s Z = ((\<exists>Y. P Y s Z) \<and> Y=w)"
@@ -341,7 +341,7 @@
subsection "RefVar"
definition
- RefVar :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr "..;" 13)
+ RefVar :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr \<open>..;\<close> 13)
where "(vf ..; P) = (\<lambda>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 = \<forall>'a. triple ('a assn) term ('a assn) **)
- ("{(1_)}/ _\<succ>/ {(1_)}" [3,65,3] 75)
+ (\<open>{(1_)}/ _\<succ>/ {(1_)}\<close> [3,65,3] 75)
type_synonym 'a triples = "'a triple set"
abbreviation
var_triple :: "['a assn, var ,'a assn] \<Rightarrow> 'a triple"
- ("{(1_)}/ _=\<succ>/ {(1_)}" [3,80,3] 75)
+ (\<open>{(1_)}/ _=\<succ>/ {(1_)}\<close> [3,80,3] 75)
where "{P} e=\<succ> {Q} == {P} In2 e\<succ> {Q}"
abbreviation
expr_triple :: "['a assn, expr ,'a assn] \<Rightarrow> 'a triple"
- ("{(1_)}/ _-\<succ>/ {(1_)}" [3,80,3] 75)
+ (\<open>{(1_)}/ _-\<succ>/ {(1_)}\<close> [3,80,3] 75)
where "{P} e-\<succ> {Q} == {P} In1l e\<succ> {Q}"
abbreviation
exprs_triple :: "['a assn, expr list ,'a assn] \<Rightarrow> 'a triple"
- ("{(1_)}/ _\<doteq>\<succ>/ {(1_)}" [3,65,3] 75)
+ (\<open>{(1_)}/ _\<doteq>\<succ>/ {(1_)}\<close> [3,65,3] 75)
where "{P} e\<doteq>\<succ> {Q} == {P} In3 e\<succ> {Q}"
abbreviation
stmt_triple :: "['a assn, stmt, 'a assn] \<Rightarrow> 'a triple"
- ("{(1_)}/ ._./ {(1_)}" [3,65,3] 75)
+ (\<open>{(1_)}/ ._./ {(1_)}\<close> [3,65,3] 75)
where "{P} .c. {Q} == {P} In1r c\<succ> {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 (\<open>{(1_)}/ _>/ {(1_)}\<close> [3,65,3]75) and
+ var_triple (\<open>{(1_)}/ _=>/ {(1_)}\<close> [3,80,3] 75) and
+ expr_triple (\<open>{(1_)}/ _->/ {(1_)}\<close> [3,80,3] 75) and
+ exprs_triple (\<open>{(1_)}/ _#>/ {(1_)}\<close> [3,65,3] 75)
lemma inj_triple: "inj (\<lambda>(P,t,Q). {P} t\<succ> {Q})"
apply (rule inj_onI)
@@ -423,11 +423,11 @@
done
definition mtriples :: "('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<Rightarrow> 'sig \<Rightarrow> expr) \<Rightarrow>
- ('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times> 'sig) set \<Rightarrow> 'a triples" ("{{(1_)}/ _-\<succ>/ {(1_)} | _}"[3,65,3,65]75) where
+ ('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times> 'sig) set \<Rightarrow> 'a triples" (\<open>{{(1_)}/ _-\<succ>/ {(1_)} | _}\<close>[3,65,3,65]75) where
"{{P} tf-\<succ> {Q} | ms} = (\<lambda>(C,sig). {Normal(P C sig)} tf C sig-\<succ> {Q C sig})`ms"
definition
- triple_valid :: "prog \<Rightarrow> nat \<Rightarrow> 'a triple \<Rightarrow> bool" ("_\<Turnstile>_:_" [61,0, 58] 57)
+ triple_valid :: "prog \<Rightarrow> nat \<Rightarrow> 'a triple \<Rightarrow> bool" (\<open>_\<Turnstile>_:_\<close> [61,0, 58] 57)
where
"G\<Turnstile>n:t =
(case t of {P} t\<succ> {Q} \<Rightarrow>
@@ -435,23 +435,23 @@
(\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> Q Y' s' Z))"
abbreviation
- triples_valid:: "prog \<Rightarrow> nat \<Rightarrow> 'a triples \<Rightarrow> bool" ("_|\<Turnstile>_:_" [61,0, 58] 57)
+ triples_valid:: "prog \<Rightarrow> nat \<Rightarrow> 'a triples \<Rightarrow> bool" (\<open>_|\<Turnstile>_:_\<close> [61,0, 58] 57)
where "G|\<Turnstile>n:ts == Ball ts (triple_valid G n)"
notation (ASCII)
- triples_valid ( "_||=_:_" [61,0, 58] 57)
+ triples_valid ( \<open>_||=_:_\<close> [61,0, 58] 57)
definition
- ax_valids :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triples \<Rightarrow> bool" ("_,_|\<Turnstile>_" [61,58,58] 57)
+ ax_valids :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triples \<Rightarrow> bool" (\<open>_,_|\<Turnstile>_\<close> [61,58,58] 57)
where "(G,A|\<Turnstile>ts) = (\<forall>n. G|\<Turnstile>n:A \<longrightarrow> G|\<Turnstile>n:ts)"
abbreviation
- ax_valid :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triple \<Rightarrow> bool" ("_,_\<Turnstile>_" [61,58,58] 57)
+ ax_valid :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triple \<Rightarrow> bool" (\<open>_,_\<Turnstile>_\<close> [61,58,58] 57)
where "G,A \<Turnstile>t == G,A|\<Turnstile>{t}"
notation (ASCII)
- ax_valid ( "_,_|=_" [61,58,58] 57)
+ ax_valid ( \<open>_,_|=_\<close> [61,58,58] 57)
lemma triple_valid_def2: "G\<Turnstile>n:{P} t\<succ> {Q} =
@@ -472,8 +472,8 @@
setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close>
inductive
- ax_derivs :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool" ("_,_|\<turnstile>_" [61,58,58] 57)
- and ax_deriv :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triple \<Rightarrow> bool" ("_,_\<turnstile>_" [61,58,58] 57)
+ ax_derivs :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool" (\<open>_,_|\<turnstile>_\<close> [61,58,58] 57)
+ and ax_deriv :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triple \<Rightarrow> bool" (\<open>_,_\<turnstile>_\<close> [61,58,58] 57)
for G :: prog
where
--- 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 \<Rightarrow> nat \<Rightarrow> 'a triple \<Rightarrow> bool" ("_\<Turnstile>_\<Colon>_"[61,0, 58] 57)
+ triple_valid2 :: "prog \<Rightarrow> nat \<Rightarrow> 'a triple \<Rightarrow> bool" (\<open>_\<Turnstile>_\<Colon>_\<close>[61,0, 58] 57)
where
"G\<Turnstile>n\<Colon>t =
(case t of {P} t\<succ> {Q} \<Rightarrow>
@@ -28,7 +28,7 @@
\<close>
definition
- ax_valids2 :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool" ("_,_|\<Turnstile>\<Colon>_" [61,58,58] 57)
+ ax_valids2 :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool" (\<open>_,_|\<Turnstile>\<Colon>_\<close> [61,58,58] 57)
where "G,A|\<Turnstile>\<Colon>ts = (\<forall>n. (\<forall>t\<in>A. G\<Turnstile>n\<Colon>t) \<longrightarrow> (\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t))"
lemma triple_valid2_def2: "G\<Turnstile>n\<Colon>{P} t\<succ> {Q} =
--- 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 \<equiv> Suc 2"
-abbreviation nat4 :: nat ("4") where "4 \<equiv> Suc 3"
+abbreviation nat3 :: nat (\<open>3\<close>) where "3 \<equiv> Suc 2"
+abbreviation nat4 :: nat (\<open>4\<close>) where "4 \<equiv> Suc 3"
(* irrefl_tranclI in Transitive_Closure.thy is more general *)
lemma irrefl_tranclI': "r\<inverse> \<inter> r\<^sup>+ = {} \<Longrightarrow> \<forall>x. (x, x) \<notin> r\<^sup>+"
@@ -144,7 +144,7 @@
subsubsection "sums"
-notation case_sum (infixr "'(+')" 80)
+notation case_sum (infixr \<open>'(+')\<close> 80)
primrec the_Inl :: "'a + 'b \<Rightarrow> 'a"
where "the_Inl (Inl a) = a"
@@ -188,12 +188,12 @@
subsubsection "quantifiers for option type"
syntax
- "_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" ("(3! _:_:/ _)" [0,0,10] 10)
- "_Oex" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" ("(3? _:_:/ _)" [0,0,10] 10)
+ "_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" (\<open>(3! _:_:/ _)\<close> [0,0,10] 10)
+ "_Oex" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" (\<open>(3? _:_:/ _)\<close> [0,0,10] 10)
syntax (symbols)
- "_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" ("(3\<forall>_\<in>_:/ _)" [0,0,10] 10)
- "_Oex" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" ("(3\<exists>_\<in>_:/ _)" [0,0,10] 10)
+ "_Oall" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" (\<open>(3\<forall>_\<in>_:/ _)\<close> [0,0,10] 10)
+ "_Oex" :: "[pttrn, 'a option, bool] \<Rightarrow> bool" (\<open>(3\<exists>_\<in>_:/ _)\<close> [0,0,10] 10)
syntax_consts
"_Oall" \<rightleftharpoons> Ball and
@@ -260,7 +260,7 @@
text \<open>list patterns -- extends pre-defined type "pttrn" used in abstractions\<close>
syntax
- "_lpttrn" :: "[pttrn, pttrn] \<Rightarrow> pttrn" ("_#/_" [901,900] 900)
+ "_lpttrn" :: "[pttrn, pttrn] \<Rightarrow> pttrn" (\<open>_#/_\<close> [901,900] 900)
syntax_consts
"_lpttrn" \<rightleftharpoons> lsplit
translations
--- 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 \<Rightarrow> st \<Rightarrow> bool" ("_\<le>|_" [71,71] 70) where
+definition gext :: "st \<Rightarrow> st \<Rightarrow> bool" (\<open>_\<le>|_\<close> [71,71] 70) where
"s\<le>|s' \<equiv> \<forall>r. \<forall>obj\<in>globs s r: \<exists>obj'\<in>globs s' r: tag obj'= tag obj"
text \<open>For the the proof of type soundness we will need the
@@ -96,7 +96,7 @@
subsubsection "value conformance"
-definition conf :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>_" [71,71,71,71] 70)
+definition conf :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" (\<open>_,_\<turnstile>_\<Colon>\<preceq>_\<close> [71,71,71,71] 70)
where "G,s\<turnstile>v\<Colon>\<preceq>T = (\<exists>T'\<in>typeof (\<lambda>a. map_option obj_ty (heap s a)) v:G\<turnstile>T'\<preceq>T)"
lemma conf_cong [simp]: "G,set_locals l s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
@@ -178,7 +178,7 @@
subsubsection "value list conformance"
definition
- lconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool" ("_,_\<turnstile>_[\<Colon>\<preceq>]_" [71,71,71,71] 70)
+ lconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool" (\<open>_,_\<turnstile>_[\<Colon>\<preceq>]_\<close> [71,71,71,71] 70)
where "G,s\<turnstile>vs[\<Colon>\<preceq>]Ts = (\<forall>n. \<forall>T\<in>Ts n: \<exists>v\<in>vs n: G,s\<turnstile>v\<Colon>\<preceq>T)"
lemma lconfD: "\<lbrakk>G,s\<turnstile>vs[\<Colon>\<preceq>]Ts; Ts n = Some T\<rbrakk> \<Longrightarrow> G,s\<turnstile>(the (vs n))\<Colon>\<preceq>T"
@@ -262,7 +262,7 @@
definition
- wlconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool" ("_,_\<turnstile>_[\<sim>\<Colon>\<preceq>]_" [71,71,71,71] 70)
+ wlconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool" (\<open>_,_\<turnstile>_[\<sim>\<Colon>\<preceq>]_\<close> [71,71,71,71] 70)
where "G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]Ts = (\<forall>n. \<forall>T\<in>Ts n: \<forall> v\<in>vs n: G,s\<turnstile>v\<Colon>\<preceq>T)"
lemma wlconfD: "\<lbrakk>G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]Ts; Ts n = Some T; vs n = Some v\<rbrakk> \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T"
@@ -341,7 +341,7 @@
subsubsection "object conformance"
definition
- oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_" [71,71,71,71] 70) where
+ oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool" (\<open>_,_\<turnstile>_\<Colon>\<preceq>\<surd>_\<close> [71,71,71,71] 70) where
"(G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r) = (G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and>
(case r of
Heap a \<Rightarrow> is_type G (obj_ty obj)
@@ -376,7 +376,7 @@
subsubsection "state conformance"
definition
- conforms :: "state \<Rightarrow> env' \<Rightarrow> bool" ("_\<Colon>\<preceq>_" [71,71] 70) where
+ conforms :: "state \<Rightarrow> env' \<Rightarrow> bool" (\<open>_\<Colon>\<preceq>_\<close> [71,71] 70) where
"xs\<Colon>\<preceq>E =
(let (G, L) = E; s = snd xs; l = locals s in
(\<forall>r. \<forall>obj\<in>globs s r: G,s\<turnstile>obj \<Colon>\<preceq>\<surd>r) \<and> G,s\<turnstile>l [\<sim>\<Colon>\<preceq>]L \<and>
--- 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\<noteq>Object \<and> (\<exists>c\<in>class G C: super c = D)}"
abbreviation
- subcls1_syntax :: "prog => [qtname, qtname] => bool" ("_\<turnstile>_\<prec>\<^sub>C1_" [71,71,71] 70)
+ subcls1_syntax :: "prog => [qtname, qtname] => bool" (\<open>_\<turnstile>_\<prec>\<^sub>C1_\<close> [71,71,71] 70)
where "G\<turnstile>C \<prec>\<^sub>C1 D == (C,D) \<in> subcls1 G"
abbreviation
- subclseq_syntax :: "prog => [qtname, qtname] => bool" ("_\<turnstile>_\<preceq>\<^sub>C _" [71,71,71] 70)
+ subclseq_syntax :: "prog => [qtname, qtname] => bool" (\<open>_\<turnstile>_\<preceq>\<^sub>C _\<close> [71,71,71] 70)
where "G\<turnstile>C \<preceq>\<^sub>C D == (C,D) \<in>(subcls1 G)\<^sup>*" (* cf. 8.1.3 *)
abbreviation
- subcls_syntax :: "prog => [qtname, qtname] => bool" ("_\<turnstile>_\<prec>\<^sub>C _" [71,71,71] 70)
+ subcls_syntax :: "prog => [qtname, qtname] => bool" (\<open>_\<turnstile>_\<prec>\<^sub>C _\<close> [71,71,71] 70)
where "G\<turnstile>C \<prec>\<^sub>C D == (C,D) \<in>(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 (\<open>_|-_<:C1_\<close> [71,71,71] 70) and
+ subclseq_syntax (\<open>_|-_<=:C _\<close>[71,71,71] 70) and
+ subcls_syntax (\<open>_|-_<:C _\<close>[71,71,71] 70)
lemma subint1I: "\<lbrakk>iface G I = Some i; J \<in> set (isuperIfs i)\<rbrakk>
\<Longrightarrow> (I,J) \<in> subint1 G"
--- 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\<close>
primrec
- accessible_in :: "prog \<Rightarrow> ty \<Rightarrow> pname \<Rightarrow> bool" ("_ \<turnstile> _ accessible'_in _" [61,61,61] 60) and
- rt_accessible_in :: "prog \<Rightarrow> ref_ty \<Rightarrow> pname \<Rightarrow> bool" ("_ \<turnstile> _ accessible'_in'' _" [61,61,61] 60)
+ accessible_in :: "prog \<Rightarrow> ty \<Rightarrow> pname \<Rightarrow> bool" (\<open>_ \<turnstile> _ accessible'_in _\<close> [61,61,61] 60) and
+ rt_accessible_in :: "prog \<Rightarrow> ref_ty \<Rightarrow> pname \<Rightarrow> bool" (\<open>_ \<turnstile> _ accessible'_in'' _\<close> [61,61,61] 60)
where
"G\<turnstile>(PrimT p) accessible_in pack = True"
| accessible_in_RefT_simp:
@@ -519,7 +519,7 @@
\end{itemize}
\<close>
definition
- inheritable_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool" ("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60)
+ inheritable_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool" (\<open>_ \<turnstile> _ inheritable'_in _\<close> [61,61,61] 60)
where
"G\<turnstile>membr inheritable_in pack =
(case (accmodi membr) of
@@ -531,13 +531,13 @@
abbreviation
Method_inheritable_in_syntax::
"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> pname \<Rightarrow> bool"
- ("_ \<turnstile>Method _ inheritable'_in _ " [61,61,61] 60)
+ (\<open>_ \<turnstile>Method _ inheritable'_in _ \<close> [61,61,61] 60)
where "G\<turnstile>Method m inheritable_in p == G\<turnstile>methdMembr m inheritable_in p"
abbreviation
Methd_inheritable_in::
"prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> pname \<Rightarrow> bool"
- ("_ \<turnstile>Methd _ _ inheritable'_in _ " [61,61,61,61] 60)
+ (\<open>_ \<turnstile>Methd _ _ inheritable'_in _ \<close> [61,61,61,61] 60)
where "G\<turnstile>Methd s m inheritable_in p == G\<turnstile>(method s m) inheritable_in p"
subsubsection "declared-in/undeclared-in"
@@ -557,7 +557,7 @@
| Some c \<Rightarrow> table_of (cfields c))"
definition
- declared_in :: "prog \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile> _ declared'_in _" [61,61,61] 60)
+ declared_in :: "prog \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool" (\<open>_\<turnstile> _ declared'_in _\<close> [61,61,61] 60)
where
"G\<turnstile>m declared_in C = (case m of
fdecl (fn,f ) \<Rightarrow> cdeclaredfield G C fn = Some f
@@ -565,12 +565,12 @@
abbreviation
method_declared_in:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_\<turnstile>Method _ declared'_in _" [61,61,61] 60)
+ (\<open>_\<turnstile>Method _ declared'_in _\<close> [61,61,61] 60)
where "G\<turnstile>Method m declared_in C == G\<turnstile>mdecl (mthd m) declared_in C"
abbreviation
methd_declared_in:: "prog \<Rightarrow> sig \<Rightarrow>(qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_\<turnstile>Methd _ _ declared'_in _" [61,61,61,61] 60)
+ (\<open>_\<turnstile>Methd _ _ declared'_in _\<close> [61,61,61,61] 60)
where "G\<turnstile>Methd s m declared_in C == G\<turnstile>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 \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile> _ undeclared'_in _" [61,61,61] 60)
+ undeclared_in :: "prog \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool" (\<open>_\<turnstile> _ undeclared'_in _\<close> [61,61,61] 60)
where
"G\<turnstile>m undeclared_in C = (case m of
fid fn \<Rightarrow> cdeclaredfield G C fn = None
@@ -595,7 +595,7 @@
inductive
members :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile> _ member'_of _" [61,61,61] 60)
+ (\<open>_ \<turnstile> _ member'_of _\<close> [61,61,61] 60)
for G :: prog
where
@@ -614,21 +614,21 @@
abbreviation
method_member_of:: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Method _ member'_of _" [61,61,61] 60)
+ (\<open>_ \<turnstile>Method _ member'_of _\<close> [61,61,61] 60)
where "G\<turnstile>Method m member_of C == G\<turnstile>(methdMembr m) member_of C"
abbreviation
methd_member_of:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Methd _ _ member'_of _" [61,61,61,61] 60)
+ (\<open>_ \<turnstile>Methd _ _ member'_of _\<close> [61,61,61,61] 60)
where "G\<turnstile>Methd s m member_of C == G\<turnstile>(method s m) member_of C"
abbreviation
fieldm_member_of:: "prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Field _ _ member'_of _" [61,61,61] 60)
+ (\<open>_ \<turnstile>Field _ _ member'_of _\<close> [61,61,61] 60)
where "G\<turnstile>Field n f member_of C == G\<turnstile>fieldm n f member_of C"
definition
- inherits :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool" ("_ \<turnstile> _ inherits _" [61,61,61] 60)
+ inherits :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool" (\<open>_ \<turnstile> _ inherits _\<close> [61,61,61] 60)
where
"G\<turnstile>C inherits m =
(G\<turnstile>m inheritable_in (pid C) \<and> G\<turnstile>memberid m undeclared_in C \<and>
@@ -639,7 +639,7 @@
definition
- member_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool" ("_ \<turnstile> _ member'_in _" [61,61,61] 60)
+ member_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool" (\<open>_ \<turnstile> _ member'_in _\<close> [61,61,61] 60)
where "G\<turnstile>m member_in C = (\<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> m member_of provC)"
text \<open>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 \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Method _ member'_in _" [61,61,61] 60)
+ (\<open>_ \<turnstile>Method _ member'_in _\<close> [61,61,61] 60)
where "G\<turnstile>Method m member_in C == G\<turnstile>(methdMembr m) member_in C"
abbreviation
methd_member_in:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Methd _ _ member'_in _" [61,61,61,61] 60)
+ (\<open>_ \<turnstile>Methd _ _ member'_in _\<close> [61,61,61,61] 60)
where "G\<turnstile>Methd s m member_in C == G\<turnstile>(method s m) member_in C"
lemma member_inD: "G\<turnstile>m member_in C
@@ -678,7 +678,7 @@
inductive
stat_overridesR :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
- ("_ \<turnstile> _ overrides\<^sub>S _" [61,61,61] 60)
+ (\<open>_ \<turnstile> _ overrides\<^sub>S _\<close> [61,61,61] 60)
for G :: prog
where
@@ -697,7 +697,7 @@
inductive
overridesR :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
- ("_ \<turnstile> _ overrides _" [61,61,61] 60)
+ (\<open>_ \<turnstile> _ overrides _\<close> [61,61,61] 60)
for G :: prog
where
@@ -716,18 +716,18 @@
abbreviation (input)
sig_stat_overrides::
"prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
- ("_,_\<turnstile> _ overrides\<^sub>S _" [61,61,61,61] 60)
+ (\<open>_,_\<turnstile> _ overrides\<^sub>S _\<close> [61,61,61,61] 60)
where "G,s\<turnstile>new overrides\<^sub>S old == G\<turnstile>(qmdecl s new) overrides\<^sub>S (qmdecl s old)"
abbreviation (input)
sig_overrides:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
- ("_,_\<turnstile> _ overrides _" [61,61,61,61] 60)
+ (\<open>_,_\<turnstile> _ overrides _\<close> [61,61,61,61] 60)
where "G,s\<turnstile>new overrides old == G\<turnstile>(qmdecl s new) overrides (qmdecl s old)"
subsubsection "Hiding"
definition
- hides :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" ("_\<turnstile> _ hides _" [61,61,61] 60)
+ hides :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" (\<open>_\<turnstile> _ hides _\<close> [61,61,61] 60)
where
"G\<turnstile>new hides old =
(is_static new \<and> msig new = msig old \<and>
@@ -738,7 +738,7 @@
abbreviation
sig_hides:: "prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> methd) \<Rightarrow> bool"
- ("_,_\<turnstile> _ hides _" [61,61,61,61] 60)
+ (\<open>_,_\<turnstile> _ hides _\<close> [61,61,61,61] 60)
where "G,s\<turnstile>new hides old == G\<turnstile>(qmdecl s new) hides (qmdecl s old)"
lemma hidesI:
@@ -788,7 +788,7 @@
subsubsection "permits access"
definition
- permits_acc :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" ("_ \<turnstile> _ in _ permits'_acc'_from _" [61,61,61,61] 60)
+ permits_acc :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" (\<open>_ \<turnstile> _ in _ permits'_acc'_from _\<close> [61,61,61,61] 60)
where
"G\<turnstile>membr in cls permits_acc_from accclass =
(case (accmodi membr) of
@@ -824,9 +824,9 @@
inductive
accessible_fromR :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
and accessible_from :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile> _ of _ accessible'_from _" [61,61,61,61] 60)
+ (\<open>_ \<turnstile> _ of _ accessible'_from _\<close> [61,61,61,61] 60)
and method_accessible_from :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Method _ of _ accessible'_from _" [61,61,61,61] 60)
+ (\<open>_ \<turnstile>Method _ of _ accessible'_from _\<close> [61,61,61,61] 60)
for G :: prog and accclass :: qtname
where
"G\<turnstile>membr of cls accessible_from accclass \<equiv> accessible_fromR G accclass membr cls"
@@ -851,7 +851,7 @@
abbreviation
methd_accessible_from::
"prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Methd _ _ of _ accessible'_from _" [61,61,61,61,61] 60)
+ (\<open>_ \<turnstile>Methd _ _ of _ accessible'_from _\<close> [61,61,61,61,61] 60)
where
"G\<turnstile>Methd s m of cls accessible_from accclass ==
G\<turnstile>(method s m) of cls accessible_from accclass"
@@ -859,7 +859,7 @@
abbreviation
field_accessible_from::
"prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Field _ _ of _ accessible'_from _" [61,61,61,61,61] 60)
+ (\<open>_ \<turnstile>Field _ _ of _ accessible'_from _\<close> [61,61,61,61,61] 60)
where
"G\<turnstile>Field fn f of C accessible_from accclass ==
G\<turnstile>(fieldm fn f) of C accessible_from accclass"
@@ -867,9 +867,9 @@
inductive
dyn_accessible_fromR :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
and dyn_accessible_from' :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile> _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
+ (\<open>_ \<turnstile> _ in _ dyn'_accessible'_from _\<close> [61,61,61,61] 60)
and method_dyn_accessible_from :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Method _ in _ dyn'_accessible'_from _" [61,61,61,61] 60)
+ (\<open>_ \<turnstile>Method _ in _ dyn'_accessible'_from _\<close> [61,61,61,61] 60)
for G :: prog and accclass :: qtname
where
"G\<turnstile>membr in C dyn_accessible_from accC \<equiv> dyn_accessible_fromR G accC membr C"
@@ -890,7 +890,7 @@
abbreviation
methd_dyn_accessible_from::
"prog \<Rightarrow> sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Methd _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60)
+ (\<open>_ \<turnstile>Methd _ _ in _ dyn'_accessible'_from _\<close> [61,61,61,61,61] 60)
where
"G\<turnstile>Methd s m in C dyn_accessible_from accC ==
G\<turnstile>(method s m) in C dyn_accessible_from accC"
@@ -898,7 +898,7 @@
abbreviation
field_dyn_accessible_from::
"prog \<Rightarrow> vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile>Field _ _ in _ dyn'_accessible'_from _" [61,61,61,61,61] 60)
+ (\<open>_ \<turnstile>Field _ _ in _ dyn'_accessible'_from _\<close> [61,61,61,61,61] 60)
where
"G\<turnstile>Field fn f in dynC dyn_accessible_from accC ==
G\<turnstile>(fieldm fn f) in dynC dyn_accessible_from accC"
--- 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 \<open>Lifting set operations to range of tables (map to a set)\<close>
definition
- union_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" ("_ \<Rightarrow>\<union> _" [67,67] 65)
+ union_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" (\<open>_ \<Rightarrow>\<union> _\<close> [67,67] 65)
where "A \<Rightarrow>\<union> B = (\<lambda> k. A k \<union> B k)"
definition
- intersect_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" ("_ \<Rightarrow>\<inter> _" [72,72] 71)
+ intersect_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" (\<open>_ \<Rightarrow>\<inter> _\<close> [72,72] 71)
where "A \<Rightarrow>\<inter> B = (\<lambda>k. A k \<inter> B k)"
definition
- all_union_ts :: "('a,'b) tables \<Rightarrow> 'b set \<Rightarrow> ('a,'b) tables" (infixl "\<Rightarrow>\<union>\<^sub>\<forall>" 40)
+ all_union_ts :: "('a,'b) tables \<Rightarrow> 'b set \<Rightarrow> ('a,'b) tables" (infixl \<open>\<Rightarrow>\<union>\<^sub>\<forall>\<close> 40)
where "(A \<Rightarrow>\<union>\<^sub>\<forall> B) = (\<lambda> k. A k \<union> B)"
subsubsection \<open>Binary union of tables\<close>
@@ -519,7 +519,7 @@
*)
definition
- range_inter_ts :: "('a,'b) tables \<Rightarrow> 'b set" ("\<Rightarrow>\<Inter>_" 80)
+ range_inter_ts :: "('a,'b) tables \<Rightarrow> 'b set" (\<open>\<Rightarrow>\<Inter>_\<close> 80)
where "\<Rightarrow>\<Inter>A = {x |x. \<forall> k. x \<in> A k}"
text \<open>
@@ -532,7 +532,7 @@
\<close>
inductive
- da :: "env \<Rightarrow> lname set \<Rightarrow> term \<Rightarrow> assigned \<Rightarrow> bool" ("_\<turnstile> _ \<guillemotright>_\<guillemotright> _" [65,65,65,65] 71)
+ da :: "env \<Rightarrow> lname set \<Rightarrow> term \<Rightarrow> assigned \<Rightarrow> bool" (\<open>_\<turnstile> _ \<guillemotright>_\<guillemotright> _\<close> [65,65,65,65] 71)
where
Skip: "Env\<turnstile> B \<guillemotright>\<langle>Skip\<rangle>\<guillemotright> \<lparr>nrm=B,brk=\<lambda> l. UNIV\<rparr>"
--- 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 @@
\<close>
abbreviation
- dummy_res :: "vals" ("\<diamondsuit>")
+ dummy_res :: "vals" (\<open>\<diamondsuit>\<close>)
where "\<diamondsuit> == In1 Unit"
abbreviation (input)
- val_inj_vals ("\<lfloor>_\<rfloor>\<^sub>e" 1000)
+ val_inj_vals (\<open>\<lfloor>_\<rfloor>\<^sub>e\<close> 1000)
where "\<lfloor>e\<rfloor>\<^sub>e == In1 e"
abbreviation (input)
- var_inj_vals ("\<lfloor>_\<rfloor>\<^sub>v" 1000)
+ var_inj_vals (\<open>\<lfloor>_\<rfloor>\<^sub>v\<close> 1000)
where "\<lfloor>v\<rfloor>\<^sub>v == In2 v"
abbreviation (input)
- lst_inj_vals ("\<lfloor>_\<rfloor>\<^sub>l" 1000)
+ lst_inj_vals (\<open>\<lfloor>_\<rfloor>\<^sub>l\<close> 1000)
where "\<lfloor>es\<rfloor>\<^sub>l == In3 es"
definition undefined3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals" where
@@ -170,7 +170,7 @@
done
definition
- fits :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60)
+ fits :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" (\<open>_,_\<turnstile>_ fits _\<close>[61,61,61,61]60)
where "G,s\<turnstile>a' fits T = ((\<exists>rt. T=RefT rt) \<longrightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T)"
lemma fits_Null [simp]: "G,s\<turnstile>Null fits T"
@@ -194,7 +194,7 @@
done
definition
- catch :: "prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60) where
+ catch :: "prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool" (\<open>_,_\<turnstile>catch _\<close>[61,61,61]60) where
"G,s\<turnstile>catch C = (\<exists>xc. abrupt s=Some (Xcpt xc) \<and>
G,store s\<turnstile>Addr (the_Loc xc) fits Class C)"
@@ -469,7 +469,7 @@
subsubsection "evaluation judgments"
inductive
- halloc :: "[prog,state,obj_tag,loc,state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>halloc _\<succ>_\<rightarrow> _"[61,61,61,61,61]60) for G::prog
+ halloc :: "[prog,state,obj_tag,loc,state]\<Rightarrow>bool" (\<open>_\<turnstile>_ \<midarrow>halloc _\<succ>_\<rightarrow> _\<close>[61,61,61,61,61]60) for G::prog
where \<comment> \<open>allocating objects on the heap, cf. 12.5\<close>
Abrupt:
@@ -481,7 +481,7 @@
\<Longrightarrow>
G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,init_obj G oi' (Heap a) s)"
-inductive sxalloc :: "[prog,state,state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>sxalloc\<rightarrow> _"[61,61,61]60) for G::prog
+inductive sxalloc :: "[prog,state,state]\<Rightarrow>bool" (\<open>_\<turnstile>_ \<midarrow>sxalloc\<rightarrow> _\<close>[61,61,61]60) for G::prog
where \<comment> \<open>allocating exception objects for
standard exceptions (other than OutOfMemory)\<close>
@@ -498,12 +498,12 @@
inductive
- eval :: "[prog,state,term,vals,state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>_\<succ>\<rightarrow> '(_, _')" [61,61,80,0,0]60)
- and exec ::"[prog,state,stmt ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<rightarrow> _" [61,61,65, 61]60)
- and evar ::"[prog,state,var ,vvar,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_=\<succ>_\<rightarrow> _"[61,61,90,61,61]60)
- and eval'::"[prog,state,expr ,val ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_-\<succ>_\<rightarrow> _"[61,61,80,61,61]60)
+ eval :: "[prog,state,term,vals,state]\<Rightarrow>bool" (\<open>_\<turnstile>_ \<midarrow>_\<succ>\<rightarrow> '(_, _')\<close> [61,61,80,0,0]60)
+ and exec ::"[prog,state,stmt ,state]\<Rightarrow>bool"(\<open>_\<turnstile>_ \<midarrow>_\<rightarrow> _\<close> [61,61,65, 61]60)
+ and evar ::"[prog,state,var ,vvar,state]\<Rightarrow>bool"(\<open>_\<turnstile>_ \<midarrow>_=\<succ>_\<rightarrow> _\<close>[61,61,90,61,61]60)
+ and eval'::"[prog,state,expr ,val ,state]\<Rightarrow>bool"(\<open>_\<turnstile>_ \<midarrow>_-\<succ>_\<rightarrow> _\<close>[61,61,80,61,61]60)
and evals::"[prog,state,expr list ,
- val list ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
+ val list ,state]\<Rightarrow>bool"(\<open>_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<rightarrow> _\<close>[61,61,61,61,61]60)
for G::prog
where
--- 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] \<Rightarrow> bool"
- ("_\<turnstile>_ \<midarrow>_\<succ>\<midarrow>_\<rightarrow> '(_, _')" [61,61,80,61,0,0] 60)
+ (\<open>_\<turnstile>_ \<midarrow>_\<succ>\<midarrow>_\<rightarrow> '(_, _')\<close> [61,61,80,61,0,0] 60)
and evarn :: "[prog, state, var, vvar, nat, state] \<Rightarrow> bool"
- ("_\<turnstile>_ \<midarrow>_=\<succ>_\<midarrow>_\<rightarrow> _" [61,61,90,61,61,61] 60)
+ (\<open>_\<turnstile>_ \<midarrow>_=\<succ>_\<midarrow>_\<rightarrow> _\<close> [61,61,90,61,61,61] 60)
and eval_n:: "[prog, state, expr, val, nat, state] \<Rightarrow> bool"
- ("_\<turnstile>_ \<midarrow>_-\<succ>_\<midarrow>_\<rightarrow> _" [61,61,80,61,61,61] 60)
+ (\<open>_\<turnstile>_ \<midarrow>_-\<succ>_\<midarrow>_\<rightarrow> _\<close> [61,61,80,61,61,61] 60)
and evalsn :: "[prog, state, expr list, val list, nat, state] \<Rightarrow> bool"
- ("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<midarrow>_\<rightarrow> _" [61,61,61,61,61,61] 60)
+ (\<open>_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<midarrow>_\<rightarrow> _\<close> [61,61,61,61,61,61] 60)
and execn :: "[prog, state, stmt, nat, state] \<Rightarrow> bool"
- ("_\<turnstile>_ \<midarrow>_\<midarrow>_\<rightarrow> _" [61,61,65, 61,61] 60)
+ (\<open>_\<turnstile>_ \<midarrow>_\<midarrow>_\<rightarrow> _\<close> [61,61,65, 61,61] 60)
for G :: prog
where
--- 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 \<Rightarrow> int \<Rightarrow> bool" ("(_/ in'_bounds _)" [50, 51] 50)
+ in_bounds :: "int \<Rightarrow> int \<Rightarrow> bool" (\<open>(_/ in'_bounds _)\<close> [50, 51] 50)
where "i in_bounds k = (0 \<le> i \<and> i < k)"
definition
@@ -306,11 +306,11 @@
subsection "update"
definition
- gupd :: "oref \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st" ("gupd'(_\<mapsto>_')" [10, 10] 1000)
+ gupd :: "oref \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st" (\<open>gupd'(_\<mapsto>_')\<close> [10, 10] 1000)
where "gupd r obj = case_st (\<lambda>g l. st (g(r\<mapsto>obj)) l)"
definition
- lupd :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" ("lupd'(_\<mapsto>_')" [10, 10] 1000)
+ lupd :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" (\<open>lupd'(_\<mapsto>_')\<close> [10, 10] 1000)
where "lupd vn v = case_st (\<lambda>g l. st g (l(vn\<mapsto>v)))"
definition
--- 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 = (\<lambda>k. \<Union>t\<in>ts. t k)"
definition overrides_t :: "('a, 'b) tables \<Rightarrow> ('a, 'b) tables \<Rightarrow> ('a, 'b) tables"
- (infixl "\<oplus>\<oplus>" 100)
+ (infixl \<open>\<oplus>\<oplus>\<close> 100)
where "s \<oplus>\<oplus> t = (\<lambda>k. if t k = {} then s k else t k)"
definition
hidings_entails :: "('a, 'b) tables \<Rightarrow> ('a, 'c) tables \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"
- ("_ hidings _ entails _" 20)
+ (\<open>_ hidings _ entails _\<close> 20)
where "(t hidings s entails R) = (\<forall>k. \<forall>x\<in>t k. \<forall>y\<in>s k. R x y)"
definition
\<comment> \<open>variant for unique table:\<close>
hiding_entails :: "('a, 'b) table \<Rightarrow> ('a, 'c) table \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"
- ("_ hiding _ entails _" 20)
+ (\<open>_ hiding _ entails _\<close> 20)
where "(t hiding s entails R) = (\<forall>k. \<forall>x\<in>t k: \<forall>y\<in>s k: R x y)"
definition
\<comment> \<open>variant for a unique table and conditional overriding:\<close>
cond_hiding_entails :: "('a, 'b) table \<Rightarrow> ('a, 'c) table
\<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> bool) \<Rightarrow> bool"
- ("_ hiding _ under _ entails _" 20)
+ (\<open>_ hiding _ under _ entails _\<close> 20)
where "(t hiding s under C entails R) = (\<forall>k. \<forall>x\<in>t k: \<forall>y\<in>s k: C x y \<longrightarrow> R x y)"
--- 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 \<comment> \<open>local variable (incl. parameters)\<close>
- | FVar qtname qtname bool expr vname ("{_,_,_}_.._"[10,10,10,85,99]90)
+ | FVar qtname qtname bool expr vname (\<open>{_,_,_}_.._\<close>[10,10,10,85,99]90)
\<comment> \<open>class field\<close>
\<comment> \<open>\<^term>\<open>{accC,statDeclC,stat}e..fn\<close>\<close>
\<comment> \<open>\<open>accC\<close>: accessing class (static class were\<close>
@@ -152,7 +152,7 @@
\<comment> \<open>\<open>stat\<close>: static or instance field?\<close>
\<comment> \<open>\<open>e\<close>: reference to object\<close>
\<comment> \<open>\<open>fn\<close>: field name\<close>
- | AVar expr expr ("_.[_]"[90,10 ]90)
+ | AVar expr expr (\<open>_.[_]\<close>[90,10 ]90)
\<comment> \<open>array component\<close>
\<comment> \<open>\<^term>\<open>e1.[e2]\<close>: e1 array reference; e2 index\<close>
| InsInitV stmt var
@@ -161,10 +161,10 @@
and expr
= NewC qtname \<comment> \<open>class instance creation\<close>
- | NewA ty expr ("New _[_]"[99,10 ]85)
+ | NewA ty expr (\<open>New _[_]\<close>[99,10 ]85)
\<comment> \<open>array creation\<close>
| Cast ty expr \<comment> \<open>type cast\<close>
- | Inst expr ref_ty ("_ InstOf _"[85,99] 85)
+ | Inst expr ref_ty (\<open>_ InstOf _\<close>[85,99] 85)
\<comment> \<open>instanceof\<close>
| Lit val \<comment> \<open>literal value, references not allowed\<close>
| UnOp unop expr \<comment> \<open>unary operation\<close>
@@ -172,11 +172,11 @@
| Super \<comment> \<open>special Super keyword\<close>
| Acc var \<comment> \<open>variable access\<close>
- | Ass var expr ("_:=_" [90,85 ]85)
+ | Ass var expr (\<open>_:=_\<close> [90,85 ]85)
\<comment> \<open>variable assign\<close>
- | Cond expr expr expr ("_ ? _ : _" [85,85,80]80) \<comment> \<open>conditional\<close>
+ | Cond expr expr expr (\<open>_ ? _ : _\<close> [85,85,80]80) \<comment> \<open>conditional\<close>
| Call qtname ref_ty inv_mode expr mname "(ty list)" "(expr list)"
- ("{_,_,_}_\<cdot>_'( {_}_')"[10,10,10,85,99,10,10]85)
+ (\<open>{_,_,_}_\<cdot>_'( {_}_')\<close>[10,10,10,85,99,10,10]85)
\<comment> \<open>method call\<close>
\<comment> \<open>\<^term>\<open>{accC,statT,mode}e\<cdot>mn({pTs}args)\<close> "\<close>
\<comment> \<open>\<open>accC\<close>: accessing class (static class were\<close>
@@ -199,20 +199,20 @@
and stmt
= Skip \<comment> \<open>empty statement\<close>
| Expr expr \<comment> \<open>expression statement\<close>
- | Lab jump stmt ("_\<bullet> _" [ 99,66]66)
+ | Lab jump stmt (\<open>_\<bullet> _\<close> [ 99,66]66)
\<comment> \<open>labeled statement; handles break\<close>
- | Comp stmt stmt ("_;; _" [ 66,65]65)
- | If' expr stmt stmt ("If'(_') _ Else _" [ 80,79,79]70)
- | Loop label expr stmt ("_\<bullet> While'(_') _" [ 99,80,79]70)
+ | Comp stmt stmt (\<open>_;; _\<close> [ 66,65]65)
+ | If' expr stmt stmt (\<open>If'(_') _ Else _\<close> [ 80,79,79]70)
+ | Loop label expr stmt (\<open>_\<bullet> While'(_') _\<close> [ 99,80,79]70)
| Jmp jump \<comment> \<open>break, continue, return\<close>
| Throw expr
- | TryC stmt qtname vname stmt ("Try _ Catch'(_ _') _" [79,99,80,79]70)
+ | TryC stmt qtname vname stmt (\<open>Try _ Catch'(_ _') _\<close> [79,99,80,79]70)
\<comment> \<open>\<^term>\<open>Try c1 Catch(C vn) c2\<close>\<close>
\<comment> \<open>\<open>c1\<close>: block were exception may be thrown\<close>
\<comment> \<open>\<open>C\<close>: execption class to catch\<close>
\<comment> \<open>\<open>vn\<close>: local name for exception used in \<open>c2\<close>\<close>
\<comment> \<open>\<open>c2\<close>: block to execute when exception is cateched\<close>
- | Fin stmt stmt ("_ Finally _" [ 79,79]70)
+ | Fin stmt stmt (\<open>_ Finally _\<close> [ 79,79]70)
| FinA abopt stmt \<comment> \<open>Save abruption of first statement\<close>
\<comment> \<open>technical term for smallstep sem.)\<close>
| Init qtname \<comment> \<open>class initialization\<close>
@@ -245,11 +245,11 @@
abbreviation this :: expr
where "this == Acc (LVar This)"
-abbreviation LAcc :: "vname \<Rightarrow> expr" ("!!")
+abbreviation LAcc :: "vname \<Rightarrow> expr" (\<open>!!\<close>)
where "!!v == Acc (LVar (EName (VNam v)))"
abbreviation
- LAss :: "vname \<Rightarrow> expr \<Rightarrow>stmt" ("_:==_" [90,85] 85)
+ LAss :: "vname \<Rightarrow> expr \<Rightarrow>stmt" (\<open>_:==_\<close> [90,85] 85)
where "v:==e == Expr (Ass (LVar (EName (VNam v))) e)"
abbreviation
@@ -274,19 +274,19 @@
\<close>
abbreviation (input)
- expr_inj_term :: "expr \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>e" 1000)
+ expr_inj_term :: "expr \<Rightarrow> term" (\<open>\<langle>_\<rangle>\<^sub>e\<close> 1000)
where "\<langle>e\<rangle>\<^sub>e == In1l e"
abbreviation (input)
- stmt_inj_term :: "stmt \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>s" 1000)
+ stmt_inj_term :: "stmt \<Rightarrow> term" (\<open>\<langle>_\<rangle>\<^sub>s\<close> 1000)
where "\<langle>c\<rangle>\<^sub>s == In1r c"
abbreviation (input)
- var_inj_term :: "var \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>v" 1000)
+ var_inj_term :: "var \<Rightarrow> term" (\<open>\<langle>_\<rangle>\<^sub>v\<close> 1000)
where "\<langle>v\<rangle>\<^sub>v == In2 v"
abbreviation (input)
- lst_inj_term :: "expr list \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>l" 1000)
+ lst_inj_term :: "expr list \<Rightarrow> term" (\<open>\<langle>_\<rangle>\<^sub>l\<close> 1000)
where "\<langle>es\<rangle>\<^sub>l == In3 es"
text \<open>It seems to be more elegant to have an overloaded injection like the
@@ -294,7 +294,7 @@
\<close>
class inj_term =
- fixes inj_term:: "'a \<Rightarrow> term" ("\<langle>_\<rangle>" 1000)
+ fixes inj_term:: "'a \<Rightarrow> term" (\<open>\<langle>_\<rangle>\<close> 1000)
text \<open>How this overloaded injections work can be seen in the theory
\<open>DefiniteAssignment\<close>. Other big inductive relations on
--- 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 \<times> state,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>1 _"[61,82,82] 81)
+ step :: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" (\<open>_\<turnstile>_ \<mapsto>1 _\<close>[61,82,82] 81)
for G :: prog
where
@@ -344,11 +344,11 @@
"G\<turnstile>(\<langle>InsInitE Skip SKIP\<rangle>,Norm s) \<mapsto>1 (\<langle>SKIP\<rangle>,Norm s)"
abbreviation
- stepn:: "[prog, term \<times> state,nat,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>_ _"[61,82,82] 81)
+ stepn:: "[prog, term \<times> state,nat,term \<times> state] \<Rightarrow> bool" (\<open>_\<turnstile>_ \<mapsto>_ _\<close>[61,82,82] 81)
where "G\<turnstile>p \<mapsto>n p' \<equiv> (p,p') \<in> {(x, y). step G x y}^^n"
abbreviation
- steptr:: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" ("_\<turnstile>_ \<mapsto>* _"[61,82,82] 81)
+ steptr:: "[prog,term \<times> state,term \<times> state] \<Rightarrow> bool" (\<open>_\<turnstile>_ \<mapsto>* _\<close>[61,82,82] 81)
where "G\<turnstile>p \<mapsto>* p' \<equiv> (p,p') \<in> {(x, y). step G x y}\<^sup>*"
(* Equivalenzen:
--- 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 \<Rightarrow> ty" ("_.[]" [90] 90)
+abbreviation Array :: "ty \<Rightarrow> ty" (\<open>_.[]\<close> [90] 90)
where "T.[] == RefT (ArrayT T)"
definition
--- 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" ("_\<turnstile>_\<prec>I1_" [71,71,71] 70)
+ subint1_syntax :: "prog => [qtname, qtname] => bool" (\<open>_\<turnstile>_\<prec>I1_\<close> [71,71,71] 70)
where "G\<turnstile>I \<prec>I1 J == (I,J) \<in> subint1 G"
abbreviation
- subint_syntax :: "prog => [qtname, qtname] => bool" ("_\<turnstile>_\<preceq>I _" [71,71,71] 70)
+ subint_syntax :: "prog => [qtname, qtname] => bool" (\<open>_\<turnstile>_\<preceq>I _\<close> [71,71,71] 70)
where "G\<turnstile>I \<preceq>I J == (I,J) \<in>(subint1 G)\<^sup>*" \<comment> \<open>cf. 9.1.3\<close>
abbreviation
- implmt1_syntax :: "prog => [qtname, qtname] => bool" ("_\<turnstile>_\<leadsto>1_" [71,71,71] 70)
+ implmt1_syntax :: "prog => [qtname, qtname] => bool" (\<open>_\<turnstile>_\<leadsto>1_\<close> [71,71,71] 70)
where "G\<turnstile>C \<leadsto>1 I == (C,I) \<in> 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 (\<open>_|-_<:I1_\<close> [71,71,71] 70) and
+ subint_syntax (\<open>_|-_<=:I _\<close>[71,71,71] 70) and
+ implmt1_syntax (\<open>_|-_~>1_\<close> [71,71,71] 70)
subsubsection "subclass and subinterface relations"
@@ -335,7 +335,7 @@
done
inductive \<comment> \<open>implementation, cf. 8.1.4\<close>
- implmt :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile>_\<leadsto>_" [71,71,71] 70)
+ implmt :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" (\<open>_\<turnstile>_\<leadsto>_\<close> [71,71,71] 70)
for G :: prog
where
direct: "G\<turnstile>C\<leadsto>1J \<Longrightarrow> G\<turnstile>C\<leadsto>J"
@@ -371,7 +371,7 @@
inductive
\<comment> \<open>widening, viz. method invocation conversion, cf. 5.3
i.e. kind of syntactic subtyping\<close>
- widen :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<preceq>_" [71,71,71] 70)
+ widen :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" (\<open>_\<turnstile>_\<preceq>_\<close> [71,71,71] 70)
for G :: prog
where
refl: "G\<turnstile>T\<preceq>T" \<comment> \<open>identity conversion, cf. 5.1.1\<close>
@@ -569,7 +569,7 @@
done
definition
- widens :: "prog \<Rightarrow> [ty list, ty list] \<Rightarrow> bool" ("_\<turnstile>_[\<preceq>]_" [71,71,71] 70)
+ widens :: "prog \<Rightarrow> [ty list, ty list] \<Rightarrow> bool" (\<open>_\<turnstile>_[\<preceq>]_\<close> [71,71,71] 70)
where "G\<turnstile>Ts[\<preceq>]Ts' = list_all2 (\<lambda>T T'. G\<turnstile>T\<preceq>T') Ts Ts'"
lemma widens_Nil [simp]: "G\<turnstile>[][\<preceq>][]"
@@ -595,7 +595,7 @@
(* more detailed than necessary for type-safety, see above rules. *)
inductive \<comment> \<open>narrowing reference conversion, cf. 5.1.5\<close>
- narrow :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<succ>_" [71,71,71] 70)
+ narrow :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" (\<open>_\<turnstile>_\<succ>_\<close> [71,71,71] 70)
for G :: prog
where
subcls: "G\<turnstile>C\<preceq>\<^sub>C D \<Longrightarrow> G\<turnstile> Class D\<succ>Class C"
@@ -646,7 +646,7 @@
subsubsection "casting relation"
inductive \<comment> \<open>casting conversion, cf. 5.5\<close>
- cast :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_\<preceq>? _" [71,71,71] 70)
+ cast :: "prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" (\<open>_\<turnstile>_\<preceq>? _\<close> [71,71,71] 70)
for G :: prog
where
widen: "G\<turnstile>S\<preceq>T \<Longrightarrow> G\<turnstile>S\<preceq>? T"
--- 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 \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env' \<Rightarrow> bool" ("_\<le>|_\<preceq>_\<Colon>\<preceq>_" [71,71,71,71] 70)
+ assign_conforms :: "st \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env' \<Rightarrow> bool" (\<open>_\<le>|_\<preceq>_\<Colon>\<preceq>_\<close> [71,71,71,71] 70)
where
"s\<le>|f\<preceq>T\<Colon>\<preceq>E =
((\<forall>s' w. Norm s'\<Colon>\<preceq>E \<longrightarrow> fst E,s'\<turnstile>w\<Colon>\<preceq>T \<longrightarrow> s\<le>|s' \<longrightarrow> assign f w (Norm s')\<Colon>\<preceq>E) \<and>
@@ -101,7 +101,7 @@
definition
- rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool" ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_" [71,71,71,71,71,71] 70)
+ rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool" (\<open>_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_\<close> [71,71,71,71,71,71] 70)
where
"G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T =
(case T of
@@ -328,7 +328,7 @@
declare fun_upd_apply [simp del]
definition
- DynT_prop :: "[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool" ("_\<turnstile>_\<rightarrow>_\<preceq>_"[71,71,71,71]70)
+ DynT_prop :: "[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool" (\<open>_\<turnstile>_\<rightarrow>_\<preceq>_\<close>[71,71,71,71]70)
where
"G\<turnstile>mode\<rightarrow>D\<preceq>t = (mode = IntVir \<longrightarrow> is_class G D \<and>
(if (\<exists>T. t=ArrayT T) then D=Object else G\<turnstile>Class D\<preceq>RefT t))"
--- 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 @@
\<close>
(* to Table *)
definition
- entails :: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool" ("_ entails _" 20)
+ entails :: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool" (\<open>_ entails _\<close> 20)
where "(t entails P) = (\<forall>k. \<forall> x \<in> t k: P x)"
lemma entailsD:
@@ -2093,7 +2093,7 @@
Array Object
\<close>
primrec valid_lookup_cls:: "prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> bool \<Rightarrow> bool"
- ("_,_ \<turnstile> _ valid'_lookup'_cls'_for _" [61,61,61,61] 60)
+ (\<open>_,_ \<turnstile> _ valid'_lookup'_cls'_for _\<close> [61,61,61,61] 60)
where
"G,NullT \<turnstile> dynC valid_lookup_cls_for static_membr = False"
| "G,IfaceT I \<turnstile> dynC valid_lookup_cls_for static_membr
--- 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 \<Rightarrow> dyn_ty \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>_" [51,51,51,51] 50)
- and wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow> stmt \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>\<surd>" [51,51,51] 50)
- and ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>-_" [51,51,51,51] 50)
- and ty_var :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>=_" [51,51,51,51] 50)
+inductive wt :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys] \<Rightarrow> bool" (\<open>_,_\<Turnstile>_\<Colon>_\<close> [51,51,51,51] 50)
+ and wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow> stmt \<Rightarrow> bool" (\<open>_,_\<Turnstile>_\<Colon>\<surd>\<close> [51,51,51] 50)
+ and ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" (\<open>_,_\<Turnstile>_\<Colon>-_\<close> [51,51,51,51] 50)
+ and ty_var :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var ,ty ] \<Rightarrow> bool" (\<open>_,_\<Turnstile>_\<Colon>=_\<close> [51,51,51,51] 50)
and ty_exprs :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list, ty list] \<Rightarrow> bool"
- ("_,_\<Turnstile>_\<Colon>\<doteq>_" [51,51,51,51] 50)
+ (\<open>_,_\<Turnstile>_\<Colon>\<doteq>_\<close> [51,51,51,51] 50)
where
"E,dt\<Turnstile>s\<Colon>\<surd> \<equiv> E,dt\<Turnstile>In1r s\<Colon>Inl (PrimT Void)"
@@ -426,31 +426,31 @@
(* for purely static typing *)
abbreviation
- wt_syntax :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>_" [51,51,51] 50)
+ wt_syntax :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" (\<open>_\<turnstile>_\<Colon>_\<close> [51,51,51] 50)
where "E\<turnstile>t\<Colon>T == E,empty_dt\<Turnstile>t\<Colon> T"
abbreviation
- wt_stmt_syntax :: "env \<Rightarrow> stmt \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<surd>" [51,51 ] 50)
+ wt_stmt_syntax :: "env \<Rightarrow> stmt \<Rightarrow> bool" (\<open>_\<turnstile>_\<Colon>\<surd>\<close> [51,51 ] 50)
where "E\<turnstile>s\<Colon>\<surd> == E\<turnstile>In1r s \<Colon> Inl (PrimT Void)"
abbreviation
- ty_expr_syntax :: "env \<Rightarrow> [expr, ty] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>-_" [51,51,51] 50)
+ ty_expr_syntax :: "env \<Rightarrow> [expr, ty] \<Rightarrow> bool" (\<open>_\<turnstile>_\<Colon>-_\<close> [51,51,51] 50)
where "E\<turnstile>e\<Colon>-T == E\<turnstile>In1l e \<Colon> Inl T"
abbreviation
- ty_var_syntax :: "env \<Rightarrow> [var, ty] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>=_" [51,51,51] 50)
+ ty_var_syntax :: "env \<Rightarrow> [var, ty] \<Rightarrow> bool" (\<open>_\<turnstile>_\<Colon>=_\<close> [51,51,51] 50)
where "E\<turnstile>e\<Colon>=T == E\<turnstile>In2 e \<Colon> Inl T"
abbreviation
- ty_exprs_syntax :: "env \<Rightarrow> [expr list, ty list] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<doteq>_" [51,51,51] 50)
+ ty_exprs_syntax :: "env \<Rightarrow> [expr list, ty list] \<Rightarrow> bool" (\<open>_\<turnstile>_\<Colon>\<doteq>_\<close> [51,51,51] 50)
where "E\<turnstile>e\<Colon>\<doteq>T == E\<turnstile>In3 e \<Colon> 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 (\<open>_|-_::_\<close> [51,51,51] 50) and
+ wt_stmt_syntax (\<open>_|-_:<>\<close> [51,51 ] 50) and
+ ty_expr_syntax (\<open>_|-_:-_\<close> [51,51,51] 50) and
+ ty_var_syntax (\<open>_|-_:=_\<close> [51,51,51] 50) and
+ ty_exprs_syntax (\<open>_|-_:#_\<close> [51,51,51] 50)
declare not_None_eq [simp del]
declare if_split [split del] if_split_asm [split del]
--- 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 (\<open>_ set[_]\<close> [22, 21] 21) =
"{A :: 'a set. |A| <o natLeq +c |UNIV :: 'k set|}"
morphisms set_bset Abs_bset
by (rule exI[of _ "{}"]) (auto simp: card_of_empty4 csum_def)
--- a/src/HOL/Cardinals/Order_Union.thy Fri Sep 20 19:07:10 2024 +0200
+++ b/src/HOL/Cardinals/Order_Union.thy Fri Sep 20 19:51:08 2024 +0200
@@ -10,10 +10,10 @@
imports Main
begin
-definition Osum :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "Osum" 60) where
+definition Osum :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix \<open>Osum\<close> 60) where
"r Osum r' = r \<union> r' \<union> {(a, a'). a \<in> Field r \<and> a' \<in> Field r'}"
-notation Osum (infix "\<union>o" 60)
+notation Osum (infix \<open>\<union>o\<close> 60)
lemma Field_Osum: "Field (r \<union>o r') = Field r \<union> Field r'"
unfolding Osum_def Field_def by blast
--- 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 \<Rightarrow> 'b rel \<Rightarrow> ('a + 'b) rel" (infixr "+o" 70)
+definition osum :: "'a rel \<Rightarrow> 'b rel \<Rightarrow> ('a + 'b) rel" (infixr \<open>+o\<close> 70)
where
"r +o r' = map_prod Inl Inl ` r \<union> map_prod Inr Inr ` r' \<union>
{(Inl a, Inr a') | a a' . a \<in> Field r \<and> a' \<in> Field r'}"
@@ -161,7 +161,7 @@
lemma map_prod_ordIso: "\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> map_prod f f ` r =o r"
by (metis dir_image_alt dir_image_ordIso ordIso_symmetric)
-definition oprod :: "'a rel \<Rightarrow> 'b rel \<Rightarrow> ('a \<times> 'b) rel" (infixr "*o" 80)
+definition oprod :: "'a rel \<Rightarrow> 'b rel \<Rightarrow> ('a \<times> 'b) rel" (infixr \<open>*o\<close> 80)
where "r *o r' = {((x1, y1), (x2, y2)).
(((y1, y2) \<in> r' - Id \<and> x1 \<in> Field r \<and> x2 \<in> Field r) \<or>
((y1, y2) \<in> Restr Id (Field r') \<and> (x1, x2) \<in> r))}"
@@ -846,7 +846,7 @@
end
-notation wo_rel2.oexp (infixl "^o" 90)
+notation wo_rel2.oexp (infixl \<open>^o\<close> 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]
--- 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 (\<open>\<^bold>[\<^bold>]\<close>) | LCons (lhd: 'a) (ltl: "'a llist") (infixr \<open>\<^bold>#\<close> 65)
subsection \<open>Finite lazy lists\<close>
code_lazy_type llist
-no_notation lazy_llist ("_")
+no_notation lazy_llist (\<open>_\<close>)
nonterminal llist_args
syntax
- "" :: "'a \<Rightarrow> llist_args" ("_")
- "_llist_args" :: "'a \<Rightarrow> llist_args \<Rightarrow> llist_args" ("_,/ _")
- "_llist" :: "llist_args => 'a list" ("\<^bold>[(_)\<^bold>]")
+ "" :: "'a \<Rightarrow> llist_args" (\<open>_\<close>)
+ "_llist_args" :: "'a \<Rightarrow> llist_args \<Rightarrow> llist_args" (\<open>_,/ _\<close>)
+ "_llist" :: "llist_args => 'a list" (\<open>\<^bold>[(_)\<^bold>]\<close>)
syntax_consts
"_llist_args" "_llist" == lazy_llist
translations
@@ -77,7 +77,7 @@
subsection \<open>Branching codatatypes\<close>
-codatatype tree = L | Node tree tree (infix "\<triangle>" 900)
+codatatype tree = L | Node tree tree (infix \<open>\<triangle>\<close> 900)
code_lazy_type tree
--- 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 "\<langle>$\<rangle>" 999)
+notation "apply" (infixl \<open>\<langle>$\<rangle>\<close> 999)
lemma bij_apply [simp]:
"bij (apply f)"
@@ -731,7 +731,7 @@
subsection \<open>Swaps\<close>
-lift_definition swap :: "'a \<Rightarrow> 'a \<Rightarrow> 'a perm" ("\<langle>_ \<leftrightarrow> _\<rangle>")
+lift_definition swap :: "'a \<Rightarrow> 'a \<Rightarrow> 'a perm" (\<open>\<langle>_ \<leftrightarrow> _\<rangle>\<close>)
is "\<lambda>a b. transpose a b"
proof
fix a b :: 'a
@@ -778,7 +778,7 @@
subsection \<open>Permutations specified by cycles\<close>
-fun cycle :: "'a list \<Rightarrow> 'a perm" ("\<langle>_\<rangle>")
+fun cycle :: "'a list \<Rightarrow> 'a perm" (\<open>\<langle>_\<rangle>\<close>)
where
"\<langle>[]\<rangle> = 1"
| "\<langle>[a]\<rangle> = 1"
@@ -794,16 +794,16 @@
bundle no_permutation_syntax
begin
- no_notation swap ("\<langle>_ \<leftrightarrow> _\<rangle>")
- no_notation cycle ("\<langle>_\<rangle>")
- no_notation "apply" (infixl "\<langle>$\<rangle>" 999)
+ no_notation swap (\<open>\<langle>_ \<leftrightarrow> _\<rangle>\<close>)
+ no_notation cycle (\<open>\<langle>_\<rangle>\<close>)
+ no_notation "apply" (infixl \<open>\<langle>$\<rangle>\<close> 999)
end
bundle permutation_syntax
begin
- notation swap ("\<langle>_ \<leftrightarrow> _\<rangle>")
- notation cycle ("\<langle>_\<rangle>")
- notation "apply" (infixl "\<langle>$\<rangle>" 999)
+ notation swap (\<open>\<langle>_ \<leftrightarrow> _\<rangle>\<close>)
+ notation cycle (\<open>\<langle>_\<rangle>\<close>)
+ notation "apply" (infixl \<open>\<langle>$\<rangle>\<close> 999)
end
unbundle no_permutation_syntax
--- 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 @@
\<close>
definition\<^marker>\<open>tag important\<close> has_contour_integral :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool"
- (infixr "has'_contour'_integral" 50)
+ (infixr \<open>has'_contour'_integral\<close> 50)
where "(f has_contour_integral i) g \<equiv>
((\<lambda>x. f(g x) * vector_derivative g (at x within {0..1}))
has_integral i) {0..1}"
definition\<^marker>\<open>tag important\<close> contour_integrable_on
- (infixr "contour'_integrable'_on" 50)
+ (infixr \<open>contour'_integrable'_on\<close> 50)
where "f contour_integrable_on g \<equiv> \<exists>i. (f has_contour_integral i) g"
definition\<^marker>\<open>tag important\<close> contour_integral
--- 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>\<open>tag important\<close>
has_laurent_expansion :: "(complex \<Rightarrow> complex) \<Rightarrow> complex fls \<Rightarrow> bool"
- (infixl "has'_laurent'_expansion" 60)
+ (infixl \<open>has'_laurent'_expansion\<close> 60)
where "(f has_laurent_expansion F) \<longleftrightarrow>
fls_conv_radius F > 0 \<and> eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
--- 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.
\<close>
definition%important meromorphic_on :: "(complex \<Rightarrow> complex) \<Rightarrow> complex set \<Rightarrow> bool"
- (infixl "(meromorphic'_on)" 50) where
+ (infixl \<open>(meromorphic'_on)\<close> 50) where
"f meromorphic_on A \<longleftrightarrow> (\<forall>z\<in>A. \<exists>F. (\<lambda>w. f (z + w)) has_laurent_expansion F)"
lemma meromorphic_at_iff: "f meromorphic_on {z} \<longleftrightarrow> isolated_singularity_at f z \<and> not_essential f z"
@@ -588,7 +588,7 @@
poles. We furthermore require that the function return 0 at any pole, for convenience.
\<close>
definition nicely_meromorphic_on :: "(complex \<Rightarrow> complex) \<Rightarrow> complex set \<Rightarrow> bool"
- (infixl "(nicely'_meromorphic'_on)" 50)
+ (infixl \<open>(nicely'_meromorphic'_on)\<close> 50)
where "f nicely_meromorphic_on A \<longleftrightarrow> f meromorphic_on A
\<and> (\<forall>z\<in>A. (is_pole f z \<and> f z=0) \<or> f \<midarrow>z\<rightarrow> f z)"
--- 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 \<open>$$\<close> 75)
lemmas fls_eqI = iffD1[OF fls_nth_inject, OF iffD2, OF fun_eq_iff, OF allI]
@@ -4601,11 +4601,11 @@
subsection \<open>Notation\<close>
-no_notation fls_nth (infixl "$$" 75)
+no_notation fls_nth (infixl \<open>$$\<close> 75)
bundle fls_notation
begin
-notation fls_nth (infixl "$$" 75)
+notation fls_nth (infixl \<open>$$\<close> 75)
end
end
\ No newline at end of file
--- 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 \<open>$\<close> 75)
lemma expand_fps_eq: "p = q \<longleftrightarrow> (\<forall>n. p $ n = q $ n)"
by (simp add: fps_nth_inject [symmetric] fun_eq_iff)
@@ -3633,7 +3633,7 @@
subsection \<open>Composition\<close>
-definition fps_compose :: "'a::semiring_1 fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps" (infixl "oo" 55)
+definition fps_compose :: "'a::semiring_1 fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps" (infixl \<open>oo\<close> 55)
where "a oo b = Abs_fps (\<lambda>n. sum (\<lambda>i. a$i * (b^i$n)) {0..n})"
lemma fps_compose_nth: "(a oo b)$n = sum (\<lambda>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 \<open>$\<close> 75)
bundle fps_notation
begin
-notation fps_nth (infixl "$" 75)
+notation fps_nth (infixl \<open>$\<close> 75)
end
end
--- 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 \<open>Auxiliary: operations for lists (later) representing coefficients\<close>
-definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "##" 65)
+definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr \<open>##\<close> 65)
where "x ## xs = (if xs = [] \<and> x = 0 then [] else x # xs)"
lemma cCons_0_Nil_eq [simp]: "0 ## [] = []"
@@ -350,9 +350,9 @@
nonterminal poly_args
syntax
- "" :: "'a \<Rightarrow> poly_args" ("_")
- "_poly_args" :: "'a \<Rightarrow> poly_args \<Rightarrow> poly_args" ("_,/ _")
- "_poly" :: "poly_args \<Rightarrow> 'a poly" ("[:(_):]")
+ "" :: "'a \<Rightarrow> poly_args" (\<open>_\<close>)
+ "_poly_args" :: "'a \<Rightarrow> poly_args \<Rightarrow> poly_args" (\<open>_,/ _\<close>)
+ "_poly" :: "poly_args \<Rightarrow> 'a poly" (\<open>[:(_):]\<close>)
syntax_consts
"_poly_args" "_poly" \<rightleftharpoons> pCons
translations
@@ -2338,7 +2338,7 @@
definition pcompose :: "'a::comm_semiring_0 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly"
where "pcompose p q = fold_coeffs (\<lambda>a c. [:a:] + q * c) p 0"
-notation pcompose (infixl "\<circ>\<^sub>p" 71)
+notation pcompose (infixl \<open>\<circ>\<^sub>p\<close> 71)
lemma pcompose_0 [simp]: "pcompose 0 q = 0"
by (simp add: pcompose_def)
@@ -6241,6 +6241,6 @@
qed (use \<open>n > 0\<close> in \<open>simp_all add: p_eq degree_power_eq\<close>)
qed
-no_notation cCons (infixr "##" 65)
+no_notation cCons (infixr \<open>##\<close> 65)
end
--- 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 \<open>Examples from the introduction\<close>
-codatatype 'a stream = SCons (shd: 'a) (stl: "'a stream") (infixr "\<lhd>" 65)
+codatatype 'a stream = SCons (shd: 'a) (stl: "'a stream") (infixr \<open>\<lhd>\<close> 65)
corec "natsFrom" :: "nat \<Rightarrow> nat stream" where
"natsFrom n = n \<lhd> natsFrom (n + 1)"
@@ -29,7 +29,7 @@
text \<open>We curry the example functions in this section because infix syntax works only for curried functions.\<close>
-corec (friend) Plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix "\<oplus>" 67) where
+corec (friend) Plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix \<open>\<oplus>\<close> 67) where
"x\<^sub>1 \<oplus> x\<^sub>2 = (shd x\<^sub>1 + shd x\<^sub>2) \<lhd> (stl x\<^sub>1 \<oplus> stl x\<^sub>2)"
section \<open>Examples from section 4\<close>
@@ -43,7 +43,7 @@
datatype 'a nelist = NEList (hd:'a) (tl:"'a list")
-primrec (transfer) snoc :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a nelist" (infix "\<rhd>" 64) where
+primrec (transfer) snoc :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a nelist" (infix \<open>\<rhd>\<close> 64) where
"[] \<rhd> a = NEList a []"
|"(b # bs) \<rhd> a = NEList b (bs @ [a])"
@@ -68,7 +68,7 @@
corec fibB :: "nat stream"
where "fibB = (0 \<lhd> 1 \<lhd> fibB) \<oplus> (0 \<lhd> fibB)"
-corec (friend) times :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix "\<otimes>" 69)
+corec (friend) times :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix \<open>\<otimes>\<close> 69)
where "xs \<otimes> ys = (shd xs * shd ys) \<lhd> xs \<otimes> stl ys \<oplus> stl xs \<otimes> ys"
corec (friend) exp :: "nat stream \<Rightarrow> nat stream"
@@ -106,7 +106,7 @@
corec catalan :: "nat \<Rightarrow> nat stream"
where "catalan n = (if n > 0 then catalan (n - 1) \<oplus> (0 \<lhd> catalan (n + 1)) else 1 \<lhd> catalan 1)"
-corec (friend) heart :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix "\<heartsuit>" 65)
+corec (friend) heart :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infix \<open>\<heartsuit>\<close> 65)
where "xs \<heartsuit> ys = SCons (shd xs * shd ys) ((((xs \<heartsuit> stl ys) \<oplus> (stl xs \<otimes> ys)) \<heartsuit> ys) \<otimes> ys)"
corec (friend) g :: "'a stream \<Rightarrow> 'a stream"
--- 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 (\<lambda>a. Put a copy)"
by (subst copy.code; simp)
-corecursive sp_comp (infixl "oo" 65) where
+corecursive sp_comp (infixl \<open>oo\<close> 65) where
"sp oo sp' = (case (out sp, out sp') of
(Put b sp, _) \<Rightarrow> In (Put b (sp oo sp'))
| (Get f, Put b sp) \<Rightarrow> In (f b) oo sp
--- 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 \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<prec>" 50)
+ fixes R :: "'a :: linorder \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open>\<prec>\<close> 50)
assumes R: "R \<in> {(<), (>)}"
shows "size {#y \<in># mset xs. y \<prec> M#} \<le> nat \<lceil>0.7 * length xs + 3\<rceil>"
proof -
@@ -406,8 +406,8 @@
text \<open>We then split that multiset into those groups whose medians is less than @{term M}
and the rest.\<close>
- define Y_small ("Y\<^sub>\<prec>") where "Y\<^sub>\<prec> = filter_mset (\<lambda>ys. median ys \<prec> M) (mset (chop 5 xs))"
- define Y_big ("Y\<^sub>\<succeq>") where "Y\<^sub>\<succeq> = filter_mset (\<lambda>ys. \<not>(median ys \<prec> M)) (mset (chop 5 xs))"
+ define Y_small (\<open>Y\<^sub>\<prec>\<close>) where "Y\<^sub>\<prec> = filter_mset (\<lambda>ys. median ys \<prec> M) (mset (chop 5 xs))"
+ define Y_big (\<open>Y\<^sub>\<succeq>\<close>) where "Y\<^sub>\<succeq> = filter_mset (\<lambda>ys. \<not>(median ys \<prec> 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>\<prec> + Y\<^sub>\<succeq>" unfolding Y_small_def Y_big_def
by (rule multiset_partition)
--- 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 \<Rightarrow> nat"
datatype 'a tree23 =
- Leaf ("\<langle>\<rangle>") |
- Node2 "'a tree23" 'a "'a tree23" ("\<langle>_, _, _\<rangle>") |
- Node3 "'a tree23" 'a "'a tree23" 'a "'a tree23" ("\<langle>_, _, _, _, _\<rangle>")
+ Leaf (\<open>\<langle>\<rangle>\<close>) |
+ Node2 "'a tree23" 'a "'a tree23" (\<open>\<langle>_, _, _\<rangle>\<close>) |
+ Node3 "'a tree23" 'a "'a tree23" 'a "'a tree23" (\<open>\<langle>_, _, _, _, _\<rangle>\<close>)
fun inorder :: "'a tree23 \<Rightarrow> 'a list" where
"inorder Leaf = []" |
--- 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 \<Rightarrow> nat"
datatype 'a tree234 =
- Leaf ("\<langle>\<rangle>") |
- Node2 "'a tree234" 'a "'a tree234" ("\<langle>_, _, _\<rangle>") |
- Node3 "'a tree234" 'a "'a tree234" 'a "'a tree234" ("\<langle>_, _, _, _, _\<rangle>") |
+ Leaf (\<open>\<langle>\<rangle>\<close>) |
+ Node2 "'a tree234" 'a "'a tree234" (\<open>\<langle>_, _, _\<rangle>\<close>) |
+ Node3 "'a tree234" 'a "'a tree234" 'a "'a tree234" (\<open>\<langle>_, _, _, _, _\<rangle>\<close>) |
Node4 "'a tree234" 'a "'a tree234" 'a "'a tree234" 'a "'a tree234"
- ("\<langle>_, _, _, _, _, _, _\<rangle>")
+ (\<open>\<langle>_, _, _, _, _, _, _\<rangle>\<close>)
fun inorder :: "'a tree234 \<Rightarrow> 'a list" where
"inorder Leaf = []" |
--- 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 \<open>+\<close> 65)
-consts Nplus :: "N \<Rightarrow> N \<Rightarrow> N" (infixl "+" 60)
+consts Nplus :: "N \<Rightarrow> N \<Rightarrow> N" (infixl \<open>+\<close> 60)
axiomatization where
Nplus_comm: "(a::N) + b = b + (a::N)"
@@ -32,7 +32,7 @@
definition par :: "dtree \<times> dtree \<Rightarrow> dtree" where
"par \<equiv> unfold par_r par_c"
-abbreviation par_abbr (infixr "\<parallel>" 80) where "tr1 \<parallel> tr2 \<equiv> par (tr1, tr2)"
+abbreviation par_abbr (infixr \<open>\<parallel>\<close> 80) where "tr1 \<parallel> tr2 \<equiv> par (tr1, tr2)"
lemma finite_par_c: "finite (par_c (tr1, tr2))"
unfolding par_c.simps apply(rule finite_UnI)
--- 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 ("\<langle>(_,/ _)\<rangle>")
+notation BNF_Def.convol (\<open>\<langle>(_,/ _)\<rangle>\<close>)
declare fset_to_fset[simp]
lemma fst_snd_convol_o[simp]: "\<langle>fst o s, snd o s\<rangle> = s"
apply(rule ext) by (simp add: convol_def)
-abbreviation sm_abbrev (infix "\<oplus>" 60)
+abbreviation sm_abbrev (infix \<open>\<oplus>\<close> 60)
where "f \<oplus> g \<equiv> Sum_Type.map_sum f g"
lemma map_sum_InlD: "(f \<oplus> g) z = Inl x \<Longrightarrow> \<exists>y. z = Inl y \<and> f y = x"
@@ -27,7 +27,7 @@
lemma map_sum_InrD: "(f \<oplus> g) z = Inr x \<Longrightarrow> \<exists>y. z = Inr y \<and> g y = x"
by (cases z) auto
-abbreviation case_sum_abbrev ("[[_,_]]" 800)
+abbreviation case_sum_abbrev (\<open>[[_,_]]\<close> 800)
where "[[f,g]] \<equiv> Sum_Type.case_sum f g"
lemma Inl_oplus_elim:
--- 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 \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<oplus>" 66) where
+primcorec plus :: "nat stream \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr \<open>\<oplus>\<close> 66) where
"shd (plus xs ys) = shd xs + shd ys"
| "stl (plus xs ys) = plus (stl xs) (stl ys)"
-definition scalar :: "nat \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr "\<cdot>" 68) where
+definition scalar :: "nat \<Rightarrow> nat stream \<Rightarrow> nat stream" (infixr \<open>\<cdot>\<close> 68) where
[simp]: "scalar n = smap (\<lambda>x. n * x)"
primcorec ones :: "nat stream" where "ones = 1 ## ones"
--- 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 (\<open>fn _ => _\<close> [0,51] 1000)
+| e_fix ExVar ExVar Ex (\<open>fix _ ( _ ) = _\<close> [0,51,51] 1000)
+| e_app Ex Ex (\<open>_ @@ _\<close> [51,51] 1000)
datatype Ty =
t_const TyConst
-| t_fun Ty Ty ("_ -> _" [51,51] 1000)
+| t_fun Ty Ty (\<open>_ -> _\<close> [51,51] 1000)
datatype 'a genClos =
- clos_mk ExVar Ex "ExVar \<rightharpoonup> 'a" ("\<langle>_ , _ , _\<rangle>" [0,0,0] 1000)
+ clos_mk ExVar Ex "ExVar \<rightharpoonup> 'a" (\<open>\<langle>_ , _ , _\<rangle>\<close> [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" (\<open>_ isof _\<close> [36,36] 50) where
isof_app: "\<lbrakk>c1 isof t1 -> t2; c2 isof t1\<rbrakk> \<Longrightarrow> c_app c1 c2 isof t2"
text \<open>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.\<close>
-inductive eval :: "[ValEnv, Ex, Val] => bool" ("_ \<turnstile> _ ---> _" [36,0,36] 50) where
+inductive eval :: "[ValEnv, Ex, Val] => bool" (\<open>_ \<turnstile> _ ---> _\<close> [36,0,36] 50) where
eval_const: "ve \<turnstile> e_const c ---> v_const c"
| eval_var2: "ev \<in> dom ve \<Longrightarrow> ve \<turnstile> e_var ev ---> the (ve ev)"
| eval_fn: "ve \<turnstile> fn ev => e ---> v_clos \<langle>ev, e, ve\<rangle>"
@@ -74,7 +74,7 @@
semantics. The relation te |- e ===> t express the expression e has the
type t in the type environment te.\<close>
-inductive elab :: "[TyEnv, Ex, Ty] => bool" ("_ \<turnstile> _ ===> _" [36,0,36] 50) where
+inductive elab :: "[TyEnv, Ex, Ty] => bool" (\<open>_ \<turnstile> _ ===> _\<close> [36,0,36] 50) where
elab_const: "c isof ty \<Longrightarrow> te \<turnstile> e_const c ===> ty"
| elab_var: "x \<in> dom te \<Longrightarrow> te \<turnstile> e_var x ===> the (te x)"
| elab_fn: "te(x \<mapsto> ty1) \<turnstile> e ===> ty2 \<Longrightarrow> te \<turnstile> 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" (\<open>_ isofenv _\<close>) where
"ve isofenv te \<equiv> (dom(ve) = dom(te) \<and>
(\<forall>x. x \<in> dom ve \<longrightarrow> (\<exists>c. the (ve x) = v_const(c) \<and> c isof the (te x))))"
-coinductive hasty :: "[Val, Ty] => bool" ("_ hasty _" [36,36] 50) where
+coinductive hasty :: "[Val, Ty] => bool" (\<open>_ hasty _\<close> [36,36] 50) where
hasty_const: "c isof t \<Longrightarrow> v_const c hasty t"
| hasty_clos: "\<lbrakk>te \<turnstile> fn ev => e ===> t; dom(ve) = dom(te) \<and>
(\<forall>x. x \<in> dom ve --> the (ve x) hasty the (te x)); cl = \<langle>ev,e,ve\<rangle>\<rbrakk> \<Longrightarrow> v_clos cl hasty t"
@@ -105,7 +105,7 @@
"v_const c hasty t"
"v_clos \<langle>xm , em , evm\<rangle> hasty u -> t"
-definition hasty_env :: "[ValEnv, TyEnv] => bool" ("_ hastyenv _ " [36,36] 35) where
+definition hasty_env :: "[ValEnv, TyEnv] => bool" (\<open>_ hastyenv _ \<close> [36,36] 35) where
[simp]: "ve hastyenv te \<equiv> (dom(ve) = dom(te) \<and>
(\<forall>x. x \<in> dom ve --> the (ve x) hasty the (te x)))"
--- 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 \<open>~\<close> 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 "(~~) \<equiv> equivclp (~)"
+abbreviation ACIcl (infix \<open>~~\<close> 64) where "(~~) \<equiv> 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 "\<approx>" 64) where
+inductive ACIEQ (infix \<open>\<approx>\<close> 64) where
a: "Alt (Alt r s) t \<approx> Alt r (Alt s t)"
| c: "Alt r s \<approx> Alt s r"
| i: "Alt r r \<approx> r"
--- 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 \<open>~\<close> 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 "(~~) \<equiv> equivclp (~)"
+abbreviation ACIDZcl (infix \<open>~~\<close> 64) where "(~~) \<equiv> 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 "\<approx>" 64) where
+inductive ACIDZEQ (infix \<open>\<approx>\<close> 64) where
a: "Alt (Alt r s) t \<approx> Alt r (Alt s t)"
| c: "Alt r s \<approx> Alt s r"
| i: "Alt r r \<approx> r"
--- 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>\<mu>_comp :: "('a, 'b, 'c) sp\<^sub>\<mu> \<Rightarrow> ('d, 'a, ('d, 'a) sp\<^sub>\<nu>) sp\<^sub>\<mu> \<Rightarrow> ('d, 'b, 'c \<times> ('d, 'a) sp\<^sub>\<nu>) sp\<^sub>\<mu>"
- (infixl "o\<^sub>\<mu>" 65)
+ (infixl \<open>o\<^sub>\<mu>\<close> 65)
where
"Put b sp o\<^sub>\<mu> fsp = Put b (sp, In fsp)"
| "Get f o\<^sub>\<mu> Put b sp = f b o\<^sub>\<mu> out sp"
@@ -70,7 +70,7 @@
termination
by (relation "lex_prod sub sub") auto
-primcorec sp\<^sub>\<nu>_comp (infixl "o\<^sub>\<nu>" 65) where
+primcorec sp\<^sub>\<nu>_comp (infixl \<open>o\<^sub>\<nu>\<close> 65) where
"out (sp o\<^sub>\<nu> sp') = map_sp\<^sub>\<mu> id (\<lambda>(sp, sp'). sp o\<^sub>\<nu> sp') (out sp o\<^sub>\<mu> out sp')"
lemma run\<^sub>\<nu>_sp\<^sub>\<nu>_comp: "run\<^sub>\<nu> (sp o\<^sub>\<nu> sp') = run\<^sub>\<nu> sp o run\<^sub>\<nu> sp'"
@@ -99,11 +99,11 @@
"sp\<^sub>\<mu>_comp2R f (Put b sp) = f b (out sp)"
| "sp\<^sub>\<mu>_comp2R f (Get h) = Get (sp\<^sub>\<mu>_comp2R f o h)"
-primrec sp\<^sub>\<mu>_comp2 (infixl "o\<^sup>*\<^sub>\<mu>" 65) where
+primrec sp\<^sub>\<mu>_comp2 (infixl \<open>o\<^sup>*\<^sub>\<mu>\<close> 65) where
"Put b sp o\<^sup>*\<^sub>\<mu> fsp = Put b (sp, In fsp)"
| "Get f o\<^sup>*\<^sub>\<mu> fsp = sp\<^sub>\<mu>_comp2R ((o\<^sup>*\<^sub>\<mu>) o f) fsp"
-primcorec sp\<^sub>\<nu>_comp2 (infixl "o\<^sup>*\<^sub>\<nu>" 65) where
+primcorec sp\<^sub>\<nu>_comp2 (infixl \<open>o\<^sup>*\<^sub>\<nu>\<close> 65) where
"out (sp o\<^sup>*\<^sub>\<nu> sp') = map_sp\<^sub>\<mu> id (\<lambda>(sp, sp'). sp o\<^sup>*\<^sub>\<nu> sp') (out sp o\<^sup>*\<^sub>\<mu> out sp')"
lemma run\<^sub>\<nu>_sp\<^sub>\<nu>_comp2: "run\<^sub>\<nu> (sp o\<^sup>*\<^sub>\<nu> sp') = run\<^sub>\<nu> sp o run\<^sub>\<nu> sp'"
@@ -175,7 +175,7 @@
bnf_axiomatization ('a, 'b) F for map: F
-notation BNF_Def.convol ("\<langle>(_,/ _)\<rangle>")
+notation BNF_Def.convol (\<open>\<langle>(_,/ _)\<rangle>\<close>)
definition \<theta> :: "('p,'a) F * 'b \<Rightarrow> ('p,'a * 'b) F" where
"\<theta> xb = F id \<langle>id, \<lambda> a. (snd xb)\<rangle> (fst xb)"
--- 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 \<Rightarrow> nat \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<^sub>\<nat>\<index>")
+definition of_natural :: "('a, 'm) ring_scheme \<Rightarrow> nat \<Rightarrow> 'a" (\<open>\<guillemotleft>_\<guillemotright>\<^sub>\<nat>\<index>\<close>)
where "\<guillemotleft>n\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> = ((\<oplus>\<^bsub>R\<^esub>) \<one>\<^bsub>R\<^esub> ^^ n) \<zero>\<^bsub>R\<^esub>"
-definition of_integer :: "('a, 'm) ring_scheme \<Rightarrow> int \<Rightarrow> 'a" ("\<guillemotleft>_\<guillemotright>\<index>")
+definition of_integer :: "('a, 'm) ring_scheme \<Rightarrow> int \<Rightarrow> 'a" (\<open>\<guillemotleft>_\<guillemotright>\<index>\<close>)
where "\<guillemotleft>i\<guillemotright>\<^bsub>R\<^esub> = (if 0 \<le> i then \<guillemotleft>nat i\<guillemotright>\<^sub>\<nat>\<^bsub>R\<^esub> else \<ominus>\<^bsub>R\<^esub> \<guillemotleft>nat (- i)\<guillemotright>\<^sub>\<nat>\<^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 \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<oslash>\<index>" 70)
+definition m_div :: "('a, 'b) ring_scheme \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<oslash>\<index>\<close> 70)
where "x \<oslash>\<^bsub>G\<^esub> y = (if y = \<zero>\<^bsub>G\<^esub> then \<zero>\<^bsub>G\<^esub> else x \<otimes>\<^bsub>G\<^esub> inv\<^bsub>G\<^esub> y)"
context field
--- 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 (\<open>Add\<close>)
+notation Minus (\<open>Minus\<close>)
+notation Mult (\<open>Mult\<close>)
+notation Inverse (\<open>Inverse\<close>)
+notation Cos (\<open>Cos\<close>)
+notation Arctan (\<open>Arctan\<close>)
+notation Abs (\<open>Abs\<close>)
+notation Max (\<open>Max\<close>)
+notation Min (\<open>Min\<close>)
+notation Pi (\<open>Pi\<close>)
+notation Sqrt (\<open>Sqrt\<close>)
+notation Exp (\<open>Exp\<close>)
+notation Powr (\<open>Powr\<close>)
+notation Ln (\<open>Ln\<close>)
+notation Power (\<open>Power\<close>)
+notation Floor (\<open>Floor\<close>)
+notation Var (\<open>Var\<close>)
+notation Num (\<open>Num\<close>)
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 (\<open>Add\<close>)
+no_notation Minus (\<open>Minus\<close>)
+no_notation Mult (\<open>Mult\<close>)
+no_notation Inverse (\<open>Inverse\<close>)
+no_notation Cos (\<open>Cos\<close>)
+no_notation Arctan (\<open>Arctan\<close>)
+no_notation Abs (\<open>Abs\<close>)
+no_notation Max (\<open>Max\<close>)
+no_notation Min (\<open>Min\<close>)
+no_notation Pi (\<open>Pi\<close>)
+no_notation Sqrt (\<open>Sqrt\<close>)
+no_notation Exp (\<open>Exp\<close>)
+no_notation Powr (\<open>Powr\<close>)
+no_notation Ln (\<open>Ln\<close>)
+no_notation Power (\<open>Power\<close>)
+no_notation Floor (\<open>Floor\<close>)
+no_notation Var (\<open>Var\<close>)
+no_notation Num (\<open>Num\<close>)
end
--- 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 \<open>Defining the basic ring operations on normalized polynomials\<close>
-function add :: "pol \<Rightarrow> pol \<Rightarrow> pol" (infixl "\<langle>+\<rangle>" 65)
+function add :: "pol \<Rightarrow> pol \<Rightarrow> pol" (infixl \<open>\<langle>+\<rangle>\<close> 65)
where
"Pc a \<langle>+\<rangle> Pc b = Pc (a + b)"
| "Pc c \<langle>+\<rangle> Pinj i P = Pinj i (P \<langle>+\<rangle> Pc c)"
@@ -132,7 +132,7 @@
by pat_completeness auto
termination by (relation "measure (\<lambda>(x, y). size x + size y)") auto
-function mul :: "pol \<Rightarrow> pol \<Rightarrow> pol" (infixl "\<langle>*\<rangle>" 70)
+function mul :: "pol \<Rightarrow> pol \<Rightarrow> pol" (infixl \<open>\<langle>*\<rangle>\<close> 70)
where
"Pc a \<langle>*\<rangle> Pc b = Pc (a * b)"
| "Pc c \<langle>*\<rangle> Pinj i P =
@@ -174,7 +174,7 @@
| "neg (PX P x Q) = PX (neg P) x (neg Q)"
text \<open>Subtraction\<close>
-definition sub :: "pol \<Rightarrow> pol \<Rightarrow> pol" (infixl "\<langle>-\<rangle>" 65)
+definition sub :: "pol \<Rightarrow> pol \<Rightarrow> pol" (infixl \<open>\<langle>-\<rangle>\<close> 65)
where "sub P Q = P \<langle>+\<rangle> neg Q"
text \<open>Square for Fast Exponentiation\<close>
--- 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 ("'(\<sqsubseteq>')") and
- less_eq ("(_/ \<sqsubseteq> _)" [51, 51] 50) and
- less ("'(\<sqsubset>')") and
- less ("(_/ \<sqsubset> _)" [51, 51] 50)
+ less_eq (\<open>'(\<sqsubseteq>')\<close>) and
+ less_eq (\<open>(_/ \<sqsubseteq> _)\<close> [51, 51] 50) and
+ less (\<open>'(\<sqsubset>')\<close>) and
+ less (\<open>(_/ \<sqsubset> _)\<close> [51, 51] 50)
end
--- 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 \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50)
+definition rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl \<open>rdvd\<close> 50)
where "x rdvd y \<longleftrightarrow> (\<exists>k::int. y = x * real_of_int k)"
lemma int_rdvd_real:
--- 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 \<equiv> C (0\<^sub>N)"
-abbreviation poly_p :: "int \<Rightarrow> poly" ("'((_)')\<^sub>p") where "(i)\<^sub>p \<equiv> C (i)\<^sub>N"
+abbreviation poly_0 :: "poly" (\<open>0\<^sub>p\<close>) where "0\<^sub>p \<equiv> C (0\<^sub>N)"
+abbreviation poly_p :: "int \<Rightarrow> poly" (\<open>'((_)')\<^sub>p\<close>) where "(i)\<^sub>p \<equiv> C (i)\<^sub>N"
subsection\<open>Boundedness, substitution and all that\<close>
@@ -116,7 +116,7 @@
declare if_cong[fundef_cong del]
declare let_cong[fundef_cong del]
-fun polyadd :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "+\<^sub>p" 60)
+fun polyadd :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl \<open>+\<^sub>p\<close> 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 \<Rightarrow> poly" ("~\<^sub>p")
+fun polyneg :: "poly \<Rightarrow> poly" (\<open>~\<^sub>p\<close>)
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 \<Rightarrow> poly \<Rightarrow> poly" (infixl "-\<^sub>p" 60)
+definition polysub :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl \<open>-\<^sub>p\<close> 60)
where "p -\<^sub>p q = polyadd p (polyneg q)"
-fun polymul :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "*\<^sub>p" 60)
+fun polymul :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl \<open>*\<^sub>p\<close> 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 \<Rightarrow> nat \<Rightarrow> poly" (infixl "^\<^sub>p" 60)
+abbreviation poly_pow :: "poly \<Rightarrow> nat \<Rightarrow> poly" (infixl \<open>^\<^sub>p\<close> 60)
where "a ^\<^sub>p k \<equiv> polypow k a"
function polynate :: "poly \<Rightarrow> 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 \<Rightarrow> 'a list \<Rightarrow>'a::{field_char_0,power}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
+abbreviation Ipoly_syntax :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{field_char_0,power}" (\<open>\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>\<close>)
where "\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<equiv> Ipoly bs p"
lemma Ipoly_CInt: "Ipoly bs (C (i, 1)) = of_int i"
--- 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 \<open>PERMUTE\<close> together with
convenient infix syntax.\<close>
-consts PERMUTE :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b" (infixr "\<bullet>" 75)
+consts PERMUTE :: "'a perm \<Rightarrow> 'b \<Rightarrow> 'b" (infixr \<open>\<bullet>\<close> 75)
text \<open>Then we add a locale for types \<^typ>\<open>'b\<close> that support
appliciation of permutations.\<close>
--- 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 \<open>Equivalence of two versions of Pappus' Axiom\<close>
no_notation
- comp (infixl "o" 55) and
- relcomp (infixr "O" 75)
+ comp (infixl \<open>o\<close> 55) and
+ relcomp (infixr \<open>O\<close> 75)
lemma p1p2:
assumes "col a b c l \<and> col d e f m"
--- 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 \<Rightarrow> 'b. cont f}"
-cpodef ('a, 'b) cfun ("(_ \<rightarrow>/ _)" [1, 0] 0) = "cfun :: ('a \<Rightarrow> 'b) set"
+cpodef ('a, 'b) cfun (\<open>(_ \<rightarrow>/ _)\<close> [1, 0] 0) = "cfun :: ('a \<Rightarrow> 'b) set"
by (auto simp: cfun_def intro: cont_const adm_cont)
type_notation (ASCII)
- cfun (infixr "->" 0)
+ cfun (infixr \<open>->\<close> 0)
notation (ASCII)
- Rep_cfun ("(_$/_)" [999,1000] 999)
+ Rep_cfun (\<open>(_$/_)\<close> [999,1000] 999)
notation
- Rep_cfun ("(_\<cdot>/_)" [999,1000] 999)
+ Rep_cfun (\<open>(_\<cdot>/_)\<close> [999,1000] 999)
subsection \<open>Syntax for continuous lambda abstraction\<close>
@@ -47,10 +47,10 @@
text \<open>Syntax for nested abstractions\<close>
syntax (ASCII)
- "_Lambda" :: "[cargs, logic] \<Rightarrow> logic" ("(3LAM _./ _)" [1000, 10] 10)
+ "_Lambda" :: "[cargs, logic] \<Rightarrow> logic" (\<open>(3LAM _./ _)\<close> [1000, 10] 10)
syntax
- "_Lambda" :: "[cargs, logic] \<Rightarrow> logic" ("(3\<Lambda> _./ _)" [1000, 10] 10)
+ "_Lambda" :: "[cargs, logic] \<Rightarrow> logic" (\<open>(3\<Lambda> _./ _)\<close> [1000, 10] 10)
syntax_consts
"_Lambda" \<rightleftharpoons> Abs_cfun
@@ -417,7 +417,7 @@
definition cfcomp :: "('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'c"
where oo_def: "cfcomp = (\<Lambda> f g x. f\<cdot>(g\<cdot>x))"
-abbreviation cfcomp_syn :: "['b \<rightarrow> 'c, 'a \<rightarrow> 'b] \<Rightarrow> 'a \<rightarrow> 'c" (infixr "oo" 100)
+abbreviation cfcomp_syn :: "['b \<rightarrow> 'c, 'a \<rightarrow> 'b] \<Rightarrow> 'a \<rightarrow> 'c" (infixr \<open>oo\<close> 100)
where "f oo g == cfcomp\<cdot>f\<cdot>g"
lemma ID1 [simp]: "ID\<cdot>x = x"
--- 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 \<open>Ideals over a preorder\<close>
locale preorder =
- fixes r :: "'a::type \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<preceq>" 50)
+ fixes r :: "'a::type \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open>\<preceq>\<close> 50)
assumes r_refl: "x \<preceq> x"
assumes r_trans: "\<lbrakk>x \<preceq> y; y \<preceq> z\<rbrakk> \<Longrightarrow> x \<preceq> z"
begin
--- 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 \<open>Basis preorder\<close>
definition
- convex_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix "\<le>\<natural>" 50) where
+ convex_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix \<open>\<le>\<natural>\<close> 50) where
"convex_le = (\<lambda>u v. u \<le>\<sharp> v \<and> u \<le>\<flat> v)"
lemma convex_le_refl [simp]: "t \<le>\<natural> t"
@@ -119,7 +119,7 @@
subsection \<open>Type definition\<close>
-typedef 'a convex_pd ("('(_')\<natural>)") =
+typedef 'a convex_pd (\<open>('(_')\<natural>)\<close>) =
"{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 \<Rightarrow> 'a convex_pd \<Rightarrow> 'a convex_pd"
- (infixl "\<union>\<natural>" 65) where
+ (infixl \<open>\<union>\<natural>\<close> 65) where
"xs \<union>\<natural> ys == convex_plus\<cdot>xs\<cdot>ys"
nonterminal convex_pd_args
syntax
- "" :: "logic \<Rightarrow> convex_pd_args" ("_")
- "_convex_pd_args" :: "logic \<Rightarrow> convex_pd_args \<Rightarrow> convex_pd_args" ("_,/ _")
- "_convex_pd" :: "convex_pd_args \<Rightarrow> logic" ("{_}\<natural>")
+ "" :: "logic \<Rightarrow> convex_pd_args" (\<open>_\<close>)
+ "_convex_pd_args" :: "logic \<Rightarrow> convex_pd_args \<Rightarrow> convex_pd_args" (\<open>_,/ _\<close>)
+ "_convex_pd" :: "convex_pd_args \<Rightarrow> logic" (\<open>{_}\<natural>\<close>)
syntax_consts
"_convex_pd_args" "_convex_pd" == convex_add
translations
@@ -350,7 +350,7 @@
syntax
"_convex_bind" :: "[logic, logic, logic] \<Rightarrow> logic"
- ("(3\<Union>\<natural>_\<in>_./ _)" [0, 0, 10] 10)
+ (\<open>(3\<Union>\<natural>_\<in>_./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_convex_bind" == convex_bind
--- 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 ("\<bullet>")
+ M = Md D | Mreq (\<open>\<bullet>\<close>)
datatype
- State = Sd D | Snil ("\<currency>")
+ State = Sd D | Snil (\<open>\<currency>\<close>)
type_synonym
SPF11 = "M fstream \<rightarrow> D fstream"
--- 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\<cdot>(flift2 (\<lambda>x. x\<in>A)))"
abbreviation
- emptystream :: "'a fstream" ("<>") where
+ emptystream :: "'a fstream" (\<open><>\<close>) where
"<> == \<bottom>"
abbreviation
- fscons' :: "'a \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_\<leadsto>_)" [66,65] 65) where
+ fscons' :: "'a \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" (\<open>(_\<leadsto>_)\<close> [66,65] 65) where
"a\<leadsto>s == fscons a\<cdot>s"
abbreviation
- fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_\<copyright>_)" [64,63] 63) where
+ fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" (\<open>(_\<copyright>_)\<close> [64,63] 63) where
"A\<copyright>s == fsfilter A\<cdot>s"
notation (ASCII)
- fscons' ("(_~>_)" [66,65] 65) and
- fsfilter' ("(_'(C')_)" [64,63] 63)
+ fscons' (\<open>(_~>_)\<close> [66,65] 65) and
+ fsfilter' (\<open>(_'(C')_)\<close> [64,63] 63)
lemma Def_maximal: "a = Def d \<Longrightarrow> a\<sqsubseteq>b \<Longrightarrow> b = Def d"
--- 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" (\<open><_>\<close> [1000] 999) where
fsingleton_def2: "fsingleton = (%a. Def a && UU)"
definition
@@ -40,15 +40,15 @@
abbreviation
- emptystream :: "'a fstream" ("<>") where
+ emptystream :: "'a fstream" (\<open><>\<close>) where
"<> == \<bottom>"
abbreviation
- fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_\<copyright>_)" [64,63] 63) where
+ fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" (\<open>(_\<copyright>_)\<close> [64,63] 63) where
"A\<copyright>s == fsfilter A\<cdot>s"
notation (ASCII)
- fsfilter' ("(_'(C')_)" [64,63] 63)
+ fsfilter' (\<open>(_'(C')_)\<close> [64,63] 63)
lemma ft_fsingleton[simp]: "ft\<cdot>(<a>) = Def a"
--- 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 \<open>Binder syntax for \<^term>\<open>fix\<close>\<close>
-abbreviation fix_syn :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" (binder "\<mu> " 10)
+abbreviation fix_syn :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" (binder \<open>\<mu> \<close> 10)
where "fix_syn (\<lambda>x. f x) \<equiv> fix\<cdot>(\<Lambda> x. f x)"
notation (ASCII)
- fix_syn (binder "FIX " 10)
+ fix_syn (binder \<open>FIX \<close> 10)
text \<open>Properties of \<^term>\<open>fix\<close>\<close>
--- 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 = (\<Lambda> m1 m2. sscase\<cdot>(\<Lambda> _. m2)\<cdot>(\<Lambda> _. m1)\<cdot>(Rep_match m1))"
abbreviation
- mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match" (infixr "+++" 65) where
+ mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match" (infixr \<open>+++\<close> 65) where
"m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
text \<open>rewrite rules for mplus\<close>
--- 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 \<Rightarrow> bool"
definition
- hoare_valid :: "[assn, com, assn] \<Rightarrow> bool" ("|= {(1_)}/ (_)/ {(1_)}" 50) where
+ hoare_valid :: "[assn, com, assn] \<Rightarrow> bool" (\<open>|= {(1_)}/ (_)/ {(1_)}\<close> 50) where
"|= {P} c {Q} = (\<forall>s t. P s \<and> D c\<cdot>(Discr s) = Def t \<longrightarrow> Q t)"
lemma WHILE_rule_sound:
--- 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 \<Rightarrow> ('a, 's) transition set"
where "trans_of = fst \<circ> snd \<circ> snd"
-abbreviation trans_of_syn ("_ \<midarrow>_\<midarrow>_\<rightarrow> _" [81, 81, 81, 81] 100)
+abbreviation trans_of_syn (\<open>_ \<midarrow>_\<midarrow>_\<rightarrow> _\<close> [81, 81, 81, 81] 100)
where "s \<midarrow>a\<midarrow>A\<rightarrow> t \<equiv> (s, a, t) \<in> trans_of A"
definition wfair_of :: "('a, 's) ioa \<Rightarrow> 'a set set"
@@ -91,7 +91,7 @@
(outputs a1 \<union> outputs a2),
(internals a1 \<union> internals a2)))"
-definition par :: "('a, 's) ioa \<Rightarrow> ('a, 't) ioa \<Rightarrow> ('a, 's * 't) ioa" (infixr "\<parallel>" 10)
+definition par :: "('a, 's) ioa \<Rightarrow> ('a, 't) ioa \<Rightarrow> ('a, 's * 't) ioa" (infixr \<open>\<parallel>\<close> 10)
where "(A \<parallel> B) =
(asig_comp (asig_of A) (asig_of B),
{pr. fst pr \<in> starts_of A \<and> snd pr \<in> starts_of B},
--- 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" (\<open>{|}\<close>)
addm :: "['a multiset, 'a] => 'a multiset"
delm :: "['a multiset, 'a] => 'a multiset"
countm :: "['a multiset, 'a => bool] => nat"
--- 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 \<Rightarrow> bool"
-definition satisfies :: "'a \<Rightarrow> 'a predicate \<Rightarrow> bool" ("_ \<Turnstile> _" [100, 9] 8)
+definition satisfies :: "'a \<Rightarrow> 'a predicate \<Rightarrow> bool" (\<open>_ \<Turnstile> _\<close> [100, 9] 8)
where "(s \<Turnstile> P) \<longleftrightarrow> P s"
-definition valid :: "'a predicate \<Rightarrow> bool" ("\<TTurnstile> _" [9] 8)
+definition valid :: "'a predicate \<Rightarrow> bool" (\<open>\<TTurnstile> _\<close> [9] 8)
where "(\<TTurnstile> P) \<longleftrightarrow> (\<forall>s. (s \<Turnstile> P))"
-definition Not :: "'a predicate \<Rightarrow> 'a predicate" ("\<^bold>\<not> _" [40] 40)
+definition Not :: "'a predicate \<Rightarrow> 'a predicate" (\<open>\<^bold>\<not> _\<close> [40] 40)
where NOT_def: "Not P s \<longleftrightarrow> \<not> P s"
-definition AND :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate" (infixr "\<^bold>\<and>" 35)
+definition AND :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate" (infixr \<open>\<^bold>\<and>\<close> 35)
where "(P \<^bold>\<and> Q) s \<longleftrightarrow> P s \<and> Q s"
-definition OR :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate" (infixr "\<^bold>\<or>" 30)
+definition OR :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate" (infixr \<open>\<^bold>\<or>\<close> 30)
where "(P \<^bold>\<or> Q) s \<longleftrightarrow> P s \<or> Q s"
-definition IMPLIES :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate" (infixr "\<^bold>\<longrightarrow>" 25)
+definition IMPLIES :: "'a predicate \<Rightarrow> 'a predicate \<Rightarrow> 'a predicate" (infixr \<open>\<^bold>\<longrightarrow>\<close> 25)
where "(P \<^bold>\<longrightarrow> Q) s \<longleftrightarrow> P s \<longrightarrow> Q s"
end
--- 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 (\<open>nil\<close>) | cons (HD :: 'a) (lazy TL :: "'a seq") (infixr \<open>##\<close> 65)
inductive Finite :: "'a seq \<Rightarrow> bool"
where
@@ -108,7 +108,7 @@
sconc_nil: "sconc \<cdot> nil \<cdot> y = y"
| sconc_cons': "x \<noteq> UU \<Longrightarrow> sconc \<cdot> (x ## xs) \<cdot> y = x ## (sconc \<cdot> xs \<cdot> y)"
-abbreviation sconc_syn :: "'a seq \<Rightarrow> 'a seq \<Rightarrow> 'a seq" (infixr "@@" 65)
+abbreviation sconc_syn :: "'a seq \<Rightarrow> 'a seq \<Rightarrow> 'a seq" (infixr \<open>@@\<close> 65)
where "xs @@ ys \<equiv> sconc \<cdot> xs \<cdot> ys"
lemma sconc_UU [simp]: "UU @@ y = UU"
--- 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 \<Rightarrow> UU
| Def y \<Rightarrow> (if P y then x ## (h \<cdot> xs) else h \<cdot> xs))))"
-abbreviation Consq_syn ("(_/\<leadsto>_)" [66, 65] 65)
+abbreviation Consq_syn (\<open>(_/\<leadsto>_)\<close> [66, 65] 65)
where "a \<leadsto> s \<equiv> Consq a \<cdot> s"
@@ -71,10 +71,10 @@
nonterminal llist_args
syntax
- "" :: "'a \<Rightarrow> llist_args" ("_")
- "_list_args" :: "'a \<Rightarrow> llist_args \<Rightarrow> llist_args" ("_,/ _")
- "_totlist" :: "llist_args \<Rightarrow> 'a Seq" ("[(_)!]")
- "_partlist" :: "llist_args \<Rightarrow> 'a Seq" ("[(_)?]")
+ "" :: "'a \<Rightarrow> llist_args" (\<open>_\<close>)
+ "_list_args" :: "'a \<Rightarrow> llist_args \<Rightarrow> llist_args" (\<open>_,/ _\<close>)
+ "_totlist" :: "llist_args \<Rightarrow> 'a Seq" (\<open>[(_)!]\<close>)
+ "_partlist" :: "llist_args \<Rightarrow> 'a Seq" (\<open>[(_)?]\<close>)
syntax_consts
"_totlist" "_partlist" \<rightleftharpoons> Consq
translations
--- 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 \<Rightarrow> bool" ("\<^bold>\<TTurnstile> _" [9] 8)
+definition validT :: "'a Seq predicate \<Rightarrow> bool" (\<open>\<^bold>\<TTurnstile> _\<close> [9] 8)
where "(\<^bold>\<TTurnstile> P) \<longleftrightarrow> (\<forall>s. s \<noteq> UU \<and> s \<noteq> nil \<longrightarrow> (s \<Turnstile> P))"
definition unlift :: "'a lift \<Rightarrow> 'a"
where "unlift x = (case x of Def y \<Rightarrow> y)"
-definition Init :: "'a predicate \<Rightarrow> 'a temporal" ("\<langle>_\<rangle>" [0] 1000)
+definition Init :: "'a predicate \<Rightarrow> 'a temporal" (\<open>\<langle>_\<rangle>\<close> [0] 1000)
where "Init P s = P (unlift (HD \<cdot> s))"
\<comment> \<open>this means that for \<open>nil\<close> and \<open>UU\<close> the effect is unpredictable\<close>
-definition Next :: "'a temporal \<Rightarrow> 'a temporal" ("\<circle>(_)" [80] 80)
+definition Next :: "'a temporal \<Rightarrow> 'a temporal" (\<open>\<circle>(_)\<close> [80] 80)
where "(\<circle>P) s \<longleftrightarrow> (if TL \<cdot> s = UU \<or> TL \<cdot> s = nil then P s else P (TL \<cdot> s))"
definition suffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool"
@@ -31,13 +31,13 @@
definition tsuffix :: "'a Seq \<Rightarrow> 'a Seq \<Rightarrow> bool"
where "tsuffix s2 s \<longleftrightarrow> s2 \<noteq> nil \<and> s2 \<noteq> UU \<and> suffix s2 s"
-definition Box :: "'a temporal \<Rightarrow> 'a temporal" ("\<box>(_)" [80] 80)
+definition Box :: "'a temporal \<Rightarrow> 'a temporal" (\<open>\<box>(_)\<close> [80] 80)
where "(\<box>P) s \<longleftrightarrow> (\<forall>s2. tsuffix s2 s \<longrightarrow> P s2)"
-definition Diamond :: "'a temporal \<Rightarrow> 'a temporal" ("\<diamond>(_)" [80] 80)
+definition Diamond :: "'a temporal \<Rightarrow> 'a temporal" (\<open>\<diamond>(_)\<close> [80] 80)
where "\<diamond>P = (\<^bold>\<not> (\<box>(\<^bold>\<not> P)))"
-definition Leadsto :: "'a temporal \<Rightarrow> 'a temporal \<Rightarrow> 'a temporal" (infixr "\<leadsto>" 22)
+definition Leadsto :: "'a temporal \<Rightarrow> 'a temporal \<Rightarrow> 'a temporal" (infixr \<open>\<leadsto>\<close> 22)
where "(P \<leadsto> Q) = (\<box>(P \<^bold>\<longrightarrow> (\<diamond>Q)))"
--- 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 \<Rightarrow> ('a option, 's) transition Seq"
where "ex2seq ex = (ex2seqC \<cdot> (mkfin (snd ex))) (fst ex)"
-definition temp_sat :: "('a, 's) execution \<Rightarrow> ('a, 's) ioa_temp \<Rightarrow> bool" (infixr "\<TTurnstile>" 22)
+definition temp_sat :: "('a, 's) execution \<Rightarrow> ('a, 's) ioa_temp \<Rightarrow> bool" (infixr \<open>\<TTurnstile>\<close> 22)
where "(ex \<TTurnstile> P) \<longleftrightarrow> ((ex2seq ex) \<Turnstile> P)"
definition validTE :: "('a, 's) ioa_temp \<Rightarrow> bool"
--- 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 \<open>Notions of implementation\<close>
-definition ioa_implements :: "('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool" (infixr "=<|" 12)
+definition ioa_implements :: "('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool" (infixr \<open>=<|\<close> 12)
where "(ioa1 =<| ioa2) \<longleftrightarrow>
(inputs (asig_of ioa1) = inputs (asig_of ioa2) \<and>
outputs (asig_of ioa1) = outputs (asig_of ioa2)) \<and>
--- 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 \<open>&&\<close> 65)
definition
smap :: "('a \<rightarrow> 'b) \<rightarrow> 'a stream \<rightarrow> 'b stream" where
@@ -22,7 +22,7 @@
If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs)"
definition
- slen :: "'a stream \<Rightarrow> enat" ("#_" [1000] 1000) where
+ slen :: "'a stream \<Rightarrow> enat" (\<open>#_\<close> [1000] 1000) where
"#s = (if stream_finite s then enat (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
@@ -37,7 +37,7 @@
"i_th = (\<lambda>i s. ft\<cdot>(i_rt i s))"
definition
- sconc :: "'a stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" (infixr "ooo" 65) where
+ sconc :: "'a stream \<Rightarrow> 'a stream \<Rightarrow> 'a stream" (infixr \<open>ooo\<close> 65) where
"s1 ooo s2 = (case #s1 of
enat n \<Rightarrow> (SOME s. (stream_take n\<cdot>s = s1) \<and> (i_rt n s = s2))
| \<infinity> \<Rightarrow> s1)"
--- 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 \<open>Further operations\<close>
definition
- flift1 :: "('a \<Rightarrow> 'b::pcpo) \<Rightarrow> ('a lift \<rightarrow> 'b)" (binder "FLIFT " 10) where
+ flift1 :: "('a \<Rightarrow> 'b::pcpo) \<Rightarrow> ('a lift \<rightarrow> 'b)" (binder \<open>FLIFT \<close> 10) where
"flift1 = (\<lambda>f. (\<Lambda> x. case_lift \<bottom> f x))"
translations
--- 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 \<open>Basis preorder\<close>
definition
- lower_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix "\<le>\<flat>" 50) where
+ lower_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix \<open>\<le>\<flat>\<close> 50) where
"lower_le = (\<lambda>u v. \<forall>x\<in>Rep_pd_basis u. \<exists>y\<in>Rep_pd_basis v. x \<sqsubseteq> y)"
lemma lower_le_refl [simp]: "t \<le>\<flat> t"
@@ -74,7 +74,7 @@
subsection \<open>Type definition\<close>
-typedef 'a lower_pd ("('(_')\<flat>)") =
+typedef 'a lower_pd (\<open>('(_')\<flat>)\<close>) =
"{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 \<Rightarrow> 'a lower_pd \<Rightarrow> 'a lower_pd"
- (infixl "\<union>\<flat>" 65) where
+ (infixl \<open>\<union>\<flat>\<close> 65) where
"xs \<union>\<flat> ys == lower_plus\<cdot>xs\<cdot>ys"
nonterminal lower_pd_args
syntax
- "" :: "logic \<Rightarrow> lower_pd_args" ("_")
- "_lower_pd_args" :: "logic \<Rightarrow> lower_pd_args \<Rightarrow> lower_pd_args" ("_,/ _")
- "_lower_pd" :: "lower_pd_args \<Rightarrow> logic" ("{_}\<flat>")
+ "" :: "logic \<Rightarrow> lower_pd_args" (\<open>_\<close>)
+ "_lower_pd_args" :: "logic \<Rightarrow> lower_pd_args \<Rightarrow> lower_pd_args" (\<open>_,/ _\<close>)
+ "_lower_pd" :: "lower_pd_args \<Rightarrow> logic" (\<open>{_}\<flat>\<close>)
syntax_consts
"_lower_pd_args" "_lower_pd" == lower_add
translations
@@ -344,7 +344,7 @@
syntax
"_lower_bind" :: "[logic, logic, logic] \<Rightarrow> logic"
- ("(3\<Union>\<flat>_\<in>_./ _)" [0, 0, 10] 10)
+ (\<open>(3\<Union>\<flat>_\<in>_./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_lower_bind" == lower_bind
--- 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: "\<exists>x. \<forall>y. x \<sqsubseteq> y"
begin
-definition bottom :: "'a" ("\<bottom>")
+definition bottom :: "'a" (\<open>\<bottom>\<close>)
where "bottom = (THE x. \<forall>y. x \<sqsubseteq> y)"
lemma minimal [iff]: "\<bottom> \<sqsubseteq> x"
--- 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 \<open><<\<close> 50)
notation
- below (infix "\<sqsubseteq>" 50)
+ below (infix \<open>\<sqsubseteq>\<close> 50)
-abbreviation not_below :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<notsqsubseteq>" 50)
+abbreviation not_below :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open>\<notsqsubseteq>\<close> 50)
where "not_below x y \<equiv> \<not> below x y"
notation (ASCII)
- not_below (infix "~<<" 50)
+ not_below (infix \<open>~<<\<close> 50)
lemma below_eq_trans: "a \<sqsubseteq> b \<Longrightarrow> b = c \<Longrightarrow> a \<sqsubseteq> c"
by (rule subst)
@@ -71,7 +71,7 @@
subsection \<open>Upper bounds\<close>
-definition is_ub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infix "<|" 55)
+definition is_ub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open><|\<close> 55)
where "S <| x \<longleftrightarrow> (\<forall>y\<in>S. y \<sqsubseteq> x)"
lemma is_ubI: "(\<And>x. x \<in> S \<Longrightarrow> x \<sqsubseteq> u) \<Longrightarrow> S <| u"
@@ -104,7 +104,7 @@
subsection \<open>Least upper bounds\<close>
-definition is_lub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infix "<<|" 55)
+definition is_lub :: "'a set \<Rightarrow> 'a \<Rightarrow> bool" (infix \<open><<|\<close> 55)
where "S <<| x \<longleftrightarrow> S <| x \<and> (\<forall>u. S <| u \<longrightarrow> x \<sqsubseteq> u)"
definition lub :: "'a set \<Rightarrow> 'a"
@@ -113,10 +113,10 @@
end
syntax (ASCII)
- "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" ("(3LUB _:_./ _)" [0,0, 10] 10)
+ "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" (\<open>(3LUB _:_./ _)\<close> [0,0, 10] 10)
syntax
- "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" ("(3\<Squnion>_\<in>_./ _)" [0,0, 10] 10)
+ "_BLub" :: "[pttrn, 'a set, 'b] \<Rightarrow> 'b" (\<open>(3\<Squnion>_\<in>_./ _)\<close> [0,0, 10] 10)
syntax_consts
"_BLub" \<rightleftharpoons> lub
@@ -127,11 +127,11 @@
context po
begin
-abbreviation Lub (binder "\<Squnion>" 10)
+abbreviation Lub (binder \<open>\<Squnion>\<close> 10)
where "\<Squnion>n. t n \<equiv> lub (range t)"
notation (ASCII)
- Lub (binder "LUB " 10)
+ Lub (binder \<open>LUB \<close> 10)
text \<open>access to some definition as inference rule\<close>
--- 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\<cdot>(liftdefl TYPE('a)) = liftemb oo liftprj"
-syntax "_LIFTDEFL" :: "type \<Rightarrow> logic" ("(1LIFTDEFL/(1'(_')))")
+syntax "_LIFTDEFL" :: "type \<Rightarrow> logic" (\<open>(1LIFTDEFL/(1'(_')))\<close>)
syntax_consts "_LIFTDEFL" \<rightleftharpoons> liftdefl
translations "LIFTDEFL('t)" \<rightleftharpoons> "CONST liftdefl TYPE('t)"
@@ -51,7 +51,7 @@
assumes liftprj_eq: "liftprj = u_map\<cdot>prj"
assumes liftdefl_eq: "liftdefl TYPE('a) = liftdefl_of\<cdot>(defl TYPE('a))"
-syntax "_DEFL" :: "type \<Rightarrow> logic" ("(1DEFL/(1'(_')))")
+syntax "_DEFL" :: "type \<Rightarrow> logic" (\<open>(1DEFL/(1'(_')))\<close>)
syntax_consts "_DEFL" \<rightleftharpoons> defl
translations "DEFL('t)" \<rightleftharpoons> "CONST defl TYPE('t)"
--- 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 "\<rightarrow>!" 0) = "{f :: 'a \<rightarrow> 'b. f\<cdot>\<bottom> = \<bottom>}"
+pcpodef ('a, 'b) sfun (infixr \<open>\<rightarrow>!\<close> 0) = "{f :: 'a \<rightarrow> 'b. f\<cdot>\<bottom> = \<bottom>}"
by simp_all
type_notation (ASCII)
- sfun (infixr "->!" 0)
+ sfun (infixr \<open>->!\<close> 0)
text \<open>TODO: Define nice syntax for abstraction, application.\<close>
--- 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 \<times> 'b. p = \<bottom> \<or> (fst p \<noteq> \<bottom> \<and> snd p \<noteq> \<bottom>)}"
-pcpodef ('a, 'b) sprod ("(_ \<otimes>/ _)" [21,20] 20) = "sprod :: ('a \<times> 'b) set"
+pcpodef ('a, 'b) sprod (\<open>(_ \<otimes>/ _)\<close> [21,20] 20) = "sprod :: ('a \<times> '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 \<open>**\<close> 20)
subsection \<open>Definitions of constants\<close>
@@ -42,9 +42,9 @@
nonterminal stuple_args
syntax
- "" :: "logic \<Rightarrow> stuple_args" ("_")
- "_stuple_args" :: "logic \<Rightarrow> stuple_args \<Rightarrow> stuple_args" ("_,/ _")
- "_stuple" :: "[logic, stuple_args] \<Rightarrow> logic" ("(1'(:_,/ _:'))")
+ "" :: "logic \<Rightarrow> stuple_args" (\<open>_\<close>)
+ "_stuple_args" :: "logic \<Rightarrow> stuple_args \<Rightarrow> stuple_args" (\<open>_,/ _\<close>)
+ "_stuple" :: "[logic, stuple_args] \<Rightarrow> logic" (\<open>(1'(:_,/ _:'))\<close>)
syntax_consts
"_stuple_args" "_stuple" \<rightleftharpoons> spair
translations
--- 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 \<and> fst (snd p) \<noteq> \<bottom> \<and> snd (snd p) = \<bottom>) \<or>
(fst p = FF \<and> fst (snd p) = \<bottom> \<and> snd (snd p) \<noteq> \<bottom>)}"
-pcpodef ('a, 'b) ssum ("(_ \<oplus>/ _)" [21, 20] 20) = "ssum :: (tr \<times> 'a \<times> 'b) set"
+pcpodef ('a, 'b) ssum (\<open>(_ \<oplus>/ _)\<close> [21, 20] 20) = "ssum :: (tr \<times> 'a \<times> '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 \<open>++\<close> 10)
subsection \<open>Definitions of constructors\<close>
--- 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 \<rightarrow> 'a \<rightarrow> tr \<rightarrow> 'a"
where "tr_case = (\<Lambda> t e (Def b). if b then t else e)"
-abbreviation cifte_syn :: "[tr, 'c, 'c] \<Rightarrow> 'c" ("(If (_)/ then (_)/ else (_))" [0, 0, 60] 60)
+abbreviation cifte_syn :: "[tr, 'c, 'c] \<Rightarrow> 'c" (\<open>(If (_)/ then (_)/ else (_))\<close> [0, 0, 60] 60)
where "If b then e1 else e2 \<equiv> tr_case\<cdot>e1\<cdot>e2\<cdot>b"
translations
@@ -82,13 +82,13 @@
definition trand :: "tr \<rightarrow> tr \<rightarrow> tr"
where andalso_def: "trand = (\<Lambda> x y. If x then y else FF)"
-abbreviation andalso_syn :: "tr \<Rightarrow> tr \<Rightarrow> tr" ("_ andalso _" [36,35] 35)
+abbreviation andalso_syn :: "tr \<Rightarrow> tr \<Rightarrow> tr" (\<open>_ andalso _\<close> [36,35] 35)
where "x andalso y \<equiv> trand\<cdot>x\<cdot>y"
definition tror :: "tr \<rightarrow> tr \<rightarrow> tr"
where orelse_def: "tror = (\<Lambda> x y. If x then TT else y)"
-abbreviation orelse_syn :: "tr \<Rightarrow> tr \<Rightarrow> tr" ("_ orelse _" [31,30] 30)
+abbreviation orelse_syn :: "tr \<Rightarrow> tr \<Rightarrow> tr" (\<open>_ orelse _\<close> [31,30] 30)
where "x orelse y \<equiv> tror\<cdot>x\<cdot>y"
definition neg :: "tr \<rightarrow> tr"
--- 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 \<open>Mixfix declarations can be given for data constructors.\<close>
-domain d5 = d5a | d5b (lazy "d5") "d5" (infixl ":#:" 70)
+domain d5 = d5a | d5b (lazy "d5") "d5" (infixl \<open>:#:\<close> 70)
lemma "d5a \<noteq> x :#: y :#: z" by simp
text \<open>Mixfix declarations can also be given for type constructors.\<close>
-domain ('a, 'b) lazypair (infixl ":*:" 25) =
- lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl ":*:" 75)
+domain ('a, 'b) lazypair (infixl \<open>:*:\<close> 25) =
+ lpair (lazy lfst :: 'a) (lazy lsnd :: 'b) (infixl \<open>:*:\<close> 75)
lemma "\<forall>p::('a :*: 'b). p \<sqsubseteq> lfst\<cdot>p :*: lsnd\<cdot>p"
by (rule allI, case_tac p, simp_all)
--- 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 \<open>choose\<close> 64)
subsection \<open>Basis for universal domain\<close>
@@ -973,6 +973,6 @@
hide_const (open) node
-notation binomial (infixl "choose" 65)
+notation binomial (infixl \<open>choose\<close> 65)
end
--- 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 \<open>Definition of new type for lifting\<close>
-datatype 'a u ("(_\<^sub>\<bottom>)" [1000] 999) = Ibottom | Iup 'a
+datatype 'a u (\<open>(_\<^sub>\<bottom>)\<close> [1000] 999) = Ibottom | Iup 'a
primrec Ifup :: "('a \<rightarrow> 'b::pcpo) \<Rightarrow> 'a u \<Rightarrow> 'b"
where
--- 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 \<open>Basis preorder\<close>
definition
- upper_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix "\<le>\<sharp>" 50) where
+ upper_le :: "'a pd_basis \<Rightarrow> 'a pd_basis \<Rightarrow> bool" (infix \<open>\<le>\<sharp>\<close> 50) where
"upper_le = (\<lambda>u v. \<forall>y\<in>Rep_pd_basis v. \<exists>x\<in>Rep_pd_basis u. x \<sqsubseteq> y)"
lemma upper_le_refl [simp]: "t \<le>\<sharp> t"
@@ -72,7 +72,7 @@
subsection \<open>Type definition\<close>
-typedef 'a upper_pd ("('(_')\<sharp>)") =
+typedef 'a upper_pd (\<open>('(_')\<sharp>)\<close>) =
"{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 \<Rightarrow> 'a upper_pd \<Rightarrow> 'a upper_pd"
- (infixl "\<union>\<sharp>" 65) where
+ (infixl \<open>\<union>\<sharp>\<close> 65) where
"xs \<union>\<sharp> ys == upper_plus\<cdot>xs\<cdot>ys"
nonterminal upper_pd_args
syntax
- "" :: "logic \<Rightarrow> upper_pd_args" ("_")
- "_upper_pd_args" :: "logic \<Rightarrow> upper_pd_args \<Rightarrow> upper_pd_args" ("_,/ _")
- "_upper_pd" :: "upper_pd_args \<Rightarrow> logic" ("{_}\<sharp>")
+ "" :: "logic \<Rightarrow> upper_pd_args" (\<open>_\<close>)
+ "_upper_pd_args" :: "logic \<Rightarrow> upper_pd_args \<Rightarrow> upper_pd_args" (\<open>_,/ _\<close>)
+ "_upper_pd" :: "upper_pd_args \<Rightarrow> logic" (\<open>{_}\<sharp>\<close>)
syntax_consts
"_upper_pd_args" "_upper_pd" == upper_add
translations
@@ -337,7 +337,7 @@
syntax
"_upper_bind" :: "[logic, logic, logic] \<Rightarrow> logic"
- ("(3\<Union>\<sharp>_\<in>_./ _)" [0, 0, 10] 10)
+ (\<open>(3\<Union>\<sharp>_\<in>_./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_upper_bind" == upper_bind
--- 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\<cdot>f\<cdot>r\<cdot>\<bottom> = \<bottom>"
by (fixrec_simp, cases r, simp_all)
-abbreviation apR (infixl "\<diamondop>" 70)
+abbreviation apR (infixl \<open>\<diamondop>\<close> 70)
where "a \<diamondop> b \<equiv> zipR\<cdot>ID\<cdot>a\<cdot>b"
text \<open>Proofs that \<open>zipR\<close> satisfies the applicative functor laws:\<close>
--- 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] \<Rightarrow> recbind" ("(2_ =/ _)" 10)
- "" :: "recbind \<Rightarrow> recbindt" ("_")
- "_recbindt" :: "[recbind, recbindt] \<Rightarrow> recbindt" ("_,/ _")
- "" :: "recbindt \<Rightarrow> recbinds" ("_")
- "_recbinds" :: "[recbindt, recbinds] \<Rightarrow> recbinds" ("_;/ _")
- "_Letrec" :: "[recbinds, logic] \<Rightarrow> logic" ("(Letrec (_)/ in (_))" 10)
+ "_recbind" :: "[logic, logic] \<Rightarrow> recbind" (\<open>(2_ =/ _)\<close> 10)
+ "" :: "recbind \<Rightarrow> recbindt" (\<open>_\<close>)
+ "_recbindt" :: "[recbind, recbindt] \<Rightarrow> recbindt" (\<open>_,/ _\<close>)
+ "" :: "recbindt \<Rightarrow> recbinds" (\<open>_\<close>)
+ "_recbinds" :: "[recbindt, recbinds] \<Rightarrow> recbinds" (\<open>_;/ _\<close>)
+ "_Letrec" :: "[recbinds, logic] \<Rightarrow> logic" (\<open>(Letrec (_)/ in (_))\<close> 10)
syntax_consts
"_recbind" "_recbinds" "_recbindt" "_Letrec" == CLetrec
--- 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 = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)"
abbreviation
- fatbar_syn :: "['a \<rightarrow> 'b match, 'a \<rightarrow> 'b match] \<Rightarrow> 'a \<rightarrow> 'b match" (infixr "\<parallel>" 60) where
+ fatbar_syn :: "['a \<rightarrow> 'b match, 'a \<rightarrow> 'b match] \<Rightarrow> 'a \<rightarrow> 'b match" (infixr \<open>\<parallel>\<close> 60) where
"m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2"
lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>"
@@ -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_ \<Rightarrow>/ _)" 10)
- "" :: "Case_syn => Cases_syn" ("_")
- "_Case2" :: "[Case_syn, Cases_syn] => Cases_syn" ("_/ | _")
- "_strip_positions" :: "'a => Case_pat" ("_")
+ "_Case_syntax":: "['a, Cases_syn] => 'b" (\<open>(Case _ of/ _)\<close> 10)
+ "_Case1" :: "[Case_pat, 'b] => Case_syn" (\<open>(2_ \<Rightarrow>/ _)\<close> 10)
+ "" :: "Case_syn => Cases_syn" (\<open>_\<close>)
+ "_Case2" :: "[Case_syn, Cases_syn] => Cases_syn" (\<open>_/ | _\<close>)
+ "_strip_positions" :: "'a => Case_pat" (\<open>_\<close>)
syntax (ASCII)
- "_Case1" :: "[Case_pat, 'b] => Case_syn" ("(2_ =>/ _)" 10)
+ "_Case1" :: "[Case_pat, 'b] => Case_syn" (\<open>(2_ =>/ _)\<close> 10)
translations
"_Case_syntax x ms" == "CONST cases\<cdot>(ms\<cdot>x)"
--- 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 \<Rightarrow> 'a" ("\<Squnion>_" [90] 90)
+definition the_lub :: "'a::order set \<Rightarrow> 'a" (\<open>\<Squnion>_\<close> [90] 90)
where "the_lub A = The (lub A)"
lemma the_lub_equality [elim?]:
--- 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 @@
\<close>
locale continuous = linearform +
- fixes norm :: "_ \<Rightarrow> real" ("\<parallel>_\<parallel>")
+ fixes norm :: "_ \<Rightarrow> real" (\<open>\<parallel>_\<parallel>\<close>)
assumes bounded: "\<exists>c. \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
declare continuous.intro [intro?] continuous_axioms.intro [intro?]
lemma continuousI [intro]:
- fixes norm :: "_ \<Rightarrow> real" ("\<parallel>_\<parallel>")
+ fixes norm :: "_ \<Rightarrow> real" (\<open>\<parallel>_\<parallel>\<close>)
assumes "linearform V f"
assumes r: "\<And>x. x \<in> V \<Longrightarrow> \<bar>f x\<bar> \<le> c * \<parallel>x\<parallel>"
shows "continuous V f norm"
@@ -69,9 +69,9 @@
\<close>
locale fn_norm =
- fixes norm :: "_ \<Rightarrow> real" ("\<parallel>_\<parallel>")
+ fixes norm :: "_ \<Rightarrow> real" (\<open>\<parallel>_\<parallel>\<close>)
fixes B defines "B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
- fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
+ fixes fn_norm (\<open>\<parallel>_\<parallel>\<hyphen>_\<close> [0, 1000] 999)
defines "\<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(B V f)"
locale normed_vectorspace_with_fn_norm = normed_vectorspace + fn_norm
--- 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 @@
\<close>
theorem norm_Hahn_Banach:
- fixes V and norm ("\<parallel>_\<parallel>")
+ fixes V and norm (\<open>\<parallel>_\<parallel>\<close>)
fixes B defines "\<And>V f. B V f \<equiv> {0} \<union> {\<bar>f x\<bar> / \<parallel>x\<parallel> | x. x \<noteq> 0 \<and> x \<in> V}"
- fixes fn_norm ("\<parallel>_\<parallel>\<hyphen>_" [0, 1000] 999)
+ fixes fn_norm (\<open>\<parallel>_\<parallel>\<hyphen>_\<close> [0, 1000] 999)
defines "\<And>V f. \<parallel>f\<parallel>\<hyphen>V \<equiv> \<Squnion>(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"
--- 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 \<Rightarrow> real" ("\<parallel>_\<parallel>")
+ fixes norm :: "'a \<Rightarrow> real" (\<open>\<parallel>_\<parallel>\<close>)
assumes ge_zero [iff?]: "x \<in> V \<Longrightarrow> 0 \<le> \<parallel>x\<parallel>"
and abs_homogenous [iff?]: "x \<in> V \<Longrightarrow> \<parallel>a \<cdot> x\<parallel> = \<bar>a\<bar> * \<parallel>x\<parallel>"
and subadditive [iff?]: "x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> \<parallel>x + y\<parallel> \<le> \<parallel>x\<parallel> + \<parallel>y\<parallel>"
--- 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 \<in> U \<Longrightarrow> a \<cdot> x \<in> U"
notation (symbols)
- subspace (infix "\<unlhd>" 50)
+ subspace (infix \<open>\<unlhd>\<close> 50)
declare vectorspace.intro [intro?] subspace.intro [intro?]
--- 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 @@
\<close>
consts
- prod :: "real \<Rightarrow> 'a::{plus,minus,zero} \<Rightarrow> 'a" (infixr "\<cdot>" 70)
+ prod :: "real \<Rightarrow> 'a::{plus,minus,zero} \<Rightarrow> 'a" (infixr \<open>\<cdot>\<close> 70)
subsection \<open>Vector space laws\<close>
--- 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 \<Rightarrow> 'b) \<Rightarrow> 'a ref \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)"
- ("_/'((_ \<rightarrow> _)')" [1000,0] 900)
+ (\<open>_/'((_ \<rightarrow> _)')\<close> [1000,0] 900)
"_fassign" :: "'a ref => id => 'v => 's com"
- ("(2_^._ :=/ _)" [70,1000,65] 61)
+ (\<open>(2_^._ :=/ _)\<close> [70,1000,65] 61)
"_faccess" :: "'a ref => ('a ref \<Rightarrow> 'v) => 'v"
- ("_^._" [65,1000] 65)
+ (\<open>_^._\<close> [65,1000] 65)
translations
"f(r \<rightarrow> v)" == "f(CONST addr r := v)"
"p^.f := e" => "f := f(p \<rightarrow> e)"
--- 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 \<Rightarrow> 'b) \<Rightarrow> 'a ref \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b)"
- ("_/'((_ \<rightarrow> _)')" [1000,0] 900)
+ (\<open>_/'((_ \<rightarrow> _)')\<close> [1000,0] 900)
"_fassign" :: "'a ref => id => 'v => 's com"
- ("(2_^._ :=/ _)" [70,1000,65] 61)
+ (\<open>(2_^._ :=/ _)\<close> [70,1000,65] 61)
"_faccess" :: "'a ref => ('a ref \<Rightarrow> 'v) => 'v"
- ("_^._" [65,1000] 65)
+ (\<open>_^._\<close> [65,1000] 65)
translations
"_refupdate f r v" == "f(CONST addr r := v)"
"p^.f := e" => "(p \<noteq> CONST Null) \<rightarrow> (f := _refupdate f p e)"
--- 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 (\<open>SKIP\<close>) where "SKIP == Basic id"
type_synonym 'a sem = "'a => 'a => bool"
--- 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 (\<open>SKIP\<close>) where "SKIP == Basic id"
type_synonym 'a sem = "'a option => 'a option => bool"
@@ -187,8 +187,8 @@
\<comment> \<open>Special syntax for guarded statements and guarded array updates:\<close>
syntax
- "_guarded_com" :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com" ("(2_ \<rightarrow>/ _)" 71)
- "_array_update" :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com" ("(2_[_] :=/ _)" [70, 65] 61)
+ "_guarded_com" :: "bool \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(2_ \<rightarrow>/ _)\<close> 71)
+ "_array_update" :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a com" (\<open>(2_[_] :=/ _)\<close> [70, 65] 61)
translations
"P \<rightarrow> c" \<rightleftharpoons> "IF P THEN c ELSE CONST Abort FI"
"a[i] := v" \<rightharpoonup> "(i < CONST length a) \<rightarrow> (a := CONST list_update a i v)"
--- 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 \<Rightarrow> 'b \<Rightarrow> 'com" ("(2_ :=/ _)" [70, 65] 61)
- "_Seq" :: "'com \<Rightarrow> 'com \<Rightarrow> 'com" ("(_;/ _)" [61,60] 60)
- "_Cond" :: "'bexp \<Rightarrow> 'com \<Rightarrow> 'com \<Rightarrow> 'com" ("(IF _/ THEN _ / ELSE _/ FI)" [0,0,0] 61)
- "_While" :: "'bexp \<Rightarrow> 'assn \<Rightarrow> 'var \<Rightarrow> 'com \<Rightarrow> 'com" ("(WHILE _/ INV {(_)} / VAR {(_)} //DO _ /OD)" [0,0,0,0] 61)
+ "_assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'com" (\<open>(2_ :=/ _)\<close> [70, 65] 61)
+ "_Seq" :: "'com \<Rightarrow> 'com \<Rightarrow> 'com" (\<open>(_;/ _)\<close> [61,60] 60)
+ "_Cond" :: "'bexp \<Rightarrow> 'com \<Rightarrow> 'com \<Rightarrow> 'com" (\<open>(IF _/ THEN _ / ELSE _/ FI)\<close> [0,0,0] 61)
+ "_While" :: "'bexp \<Rightarrow> 'assn \<Rightarrow> 'var \<Rightarrow> 'com \<Rightarrow> 'com" (\<open>(WHILE _/ INV {(_)} / VAR {(_)} //DO _ /OD)\<close> [0,0,0,0] 61)
text \<open>The \<open>VAR {_}\<close> syntax supports two variants:
\<^item> \<open>VAR {x = t}\<close> where \<open>t::nat\<close> is the decreasing expression,
@@ -23,17 +23,17 @@
\<close>
syntax
- "_While0" :: "'bexp \<Rightarrow> 'assn \<Rightarrow> 'com \<Rightarrow> 'com" ("(1WHILE _/ INV {_} //DO _ /OD)" [0,0,0] 61)
+ "_While0" :: "'bexp \<Rightarrow> 'assn \<Rightarrow> 'com \<Rightarrow> 'com" (\<open>(1WHILE _/ INV {_} //DO _ /OD)\<close> [0,0,0] 61)
text \<open>The \<open>_While0\<close> syntax is translated into the \<open>_While\<close> syntax with the trivial variant 0.
This is ok because partial correctness proofs do not make use of the variant.\<close>
syntax
- "_hoare_vars" :: "[idts, 'assn,'com, 'assn] \<Rightarrow> bool" ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
- "_hoare_vars_tc" :: "[idts, 'assn, 'com, 'assn] \<Rightarrow> bool" ("VARS _// [_] // _ // [_]" [0,0,55,0] 50)
+ "_hoare_vars" :: "[idts, 'assn,'com, 'assn] \<Rightarrow> bool" (\<open>VARS _// {_} // _ // {_}\<close> [0,0,55,0] 50)
+ "_hoare_vars_tc" :: "[idts, 'assn, 'com, 'assn] \<Rightarrow> bool" (\<open>VARS _// [_] // _ // [_]\<close> [0,0,55,0] 50)
syntax (output)
- "_hoare" :: "['assn, 'com, 'assn] \<Rightarrow> bool" ("({_}//_//{_})" [0,55,0] 50)
- "_hoare_tc" :: "['assn, 'com, 'assn] \<Rightarrow> bool" ("([_]//_//[_])" [0,55,0] 50)
+ "_hoare" :: "['assn, 'com, 'assn] \<Rightarrow> bool" (\<open>({_}//_//{_})\<close> [0,55,0] 50)
+ "_hoare_tc" :: "['assn, 'com, 'assn] \<Rightarrow> bool" (\<open>([_]//_//[_])\<close> [0,55,0] 50)
text \<open>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 \<open>x\<close> in \<open>VAR {x = t}\<close>.
--- 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)
+ (\<open>(2_^._ :=/ _)\<close> [70,1000,65] 61)
"_faccess" :: "'a::ref => ('a::ref \<Rightarrow> 'v) => 'v"
- ("_^._" [65,1000] 65)
+ (\<open>_^._\<close> [65,1000] 65)
translations
"p^.f := e" => "f := CONST fun_upd f p e"
"p^.f" => "f p"
--- 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
\<comment> \<open>Restriction of a relation\<close>
- restr ::"('a \<times> 'a) set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'a) set" ("(_/ | _)" [50, 51] 50)
+ restr ::"('a \<times> 'a) set \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a \<times> 'a) set" (\<open>(_/ | _)\<close> [50, 51] 50)
where "restr r m = {(x,y). (x,y) \<in> r \<and> \<not> m x}"
text \<open>Rewrite rules for the restriction of a relation\<close>
--- 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\<open>The semantic definition of a few connectives:\<close>
-definition ortho :: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix "\<bottom>" 55)
+definition ortho :: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix \<open>\<bottom>\<close> 55)
where "h1 \<bottom> h2 \<longleftrightarrow> dom h1 \<inter> dom h2 = {}"
definition is_empty :: "heap \<Rightarrow> bool"
@@ -51,10 +51,10 @@
bound Hs - otherwise they may bind the implicit H.\<close>
syntax
- "_emp" :: "bool" ("emp")
- "_singl" :: "nat \<Rightarrow> nat \<Rightarrow> bool" ("[_ \<mapsto> _]")
- "_star" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "**" 60)
- "_wand" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl "-*" 60)
+ "_emp" :: "bool" (\<open>emp\<close>)
+ "_singl" :: "nat \<Rightarrow> nat \<Rightarrow> bool" (\<open>[_ \<mapsto> _]\<close>)
+ "_star" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl \<open>**\<close> 60)
+ "_wand" :: "bool \<Rightarrow> bool \<Rightarrow> bool" (infixl \<open>-*\<close> 60)
(* FIXME does not handle "_idtdummy" *)
ML \<open>
--- 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 \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>- _//_//_)" [90,55,90] 50)
- and ann_hoare :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(2\<turnstile> _// _)" [60,90] 45)
+ oghoare :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" (\<open>(3\<parallel>- _//_//_)\<close> [90,55,90] 50)
+ and ann_hoare :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" (\<open>(2\<turnstile> _// _)\<close> [60,90] 45)
where
AnnBasic: "r \<subseteq> {s. f s \<in> q} \<Longrightarrow> \<turnstile> (AnnBasic r f) q"
--- 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\<open>Syntax for commands and for assertions and boolean expressions in
commands \<open>com\<close> and annotated commands \<open>ann_com\<close>.\<close>
-abbreviation Skip :: "'a com" ("SKIP" 63)
+abbreviation Skip :: "'a com" (\<open>SKIP\<close> 63)
where "SKIP \<equiv> Basic id"
-abbreviation AnnSkip :: "'a assn \<Rightarrow> 'a ann_com" ("_//SKIP" [90] 63)
+abbreviation AnnSkip :: "'a assn \<Rightarrow> 'a ann_com" (\<open>_//SKIP\<close> [90] 63)
where "r SKIP \<equiv> AnnBasic r id"
notation
- Seq ("_,,/ _" [55, 56] 55) and
- AnnSeq ("_;;/ _" [60,61] 60)
+ Seq (\<open>_,,/ _\<close> [55, 56] 55) and
+ AnnSeq (\<open>_;;/ _\<close> [60,61] 60)
syntax
- "_Assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com" ("(\<acute>_ :=/ _)" [70, 65] 61)
- "_AnnAssign" :: "'a assn \<Rightarrow> idt \<Rightarrow> 'b \<Rightarrow> 'a com" ("(_ \<acute>_ :=/ _)" [90,70,65] 61)
+ "_Assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com" (\<open>(\<acute>_ :=/ _)\<close> [70, 65] 61)
+ "_AnnAssign" :: "'a assn \<Rightarrow> idt \<Rightarrow> 'b \<Rightarrow> 'a com" (\<open>(_ \<acute>_ :=/ _)\<close> [90,70,65] 61)
translations
"\<acute>x := a" \<rightharpoonup> "CONST Basic \<guillemotleft>\<acute>(_update_name x (\<lambda>_. a))\<guillemotright>"
@@ -25,23 +25,23 @@
syntax
"_AnnCond1" :: "'a assn \<Rightarrow> 'a bexp \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"
- ("_ //IF _ /THEN _ /ELSE _ /FI" [90,0,0,0] 61)
+ (\<open>_ //IF _ /THEN _ /ELSE _ /FI\<close> [90,0,0,0] 61)
"_AnnCond2" :: "'a assn \<Rightarrow> 'a bexp \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"
- ("_ //IF _ /THEN _ /FI" [90,0,0] 61)
+ (\<open>_ //IF _ /THEN _ /FI\<close> [90,0,0] 61)
"_AnnWhile" :: "'a assn \<Rightarrow> 'a bexp \<Rightarrow> 'a assn \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"
- ("_ //WHILE _ /INV _ //DO _//OD" [90,0,0,0] 61)
+ (\<open>_ //WHILE _ /INV _ //DO _//OD\<close> [90,0,0,0] 61)
"_AnnAwait" :: "'a assn \<Rightarrow> 'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a ann_com"
- ("_ //AWAIT _ /THEN /_ /END" [90,0,0] 61)
- "_AnnAtom" :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a ann_com" ("_//\<langle>_\<rangle>" [90,0] 61)
- "_AnnWait" :: "'a assn \<Rightarrow> 'a bexp \<Rightarrow> 'a ann_com" ("_//WAIT _ END" [90,0] 61)
+ (\<open>_ //AWAIT _ /THEN /_ /END\<close> [90,0,0] 61)
+ "_AnnAtom" :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a ann_com" (\<open>_//\<langle>_\<rangle>\<close> [90,0] 61)
+ "_AnnWait" :: "'a assn \<Rightarrow> 'a bexp \<Rightarrow> 'a ann_com" (\<open>_//WAIT _ END\<close> [90,0] 61)
"_Cond" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"
- ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61)
- "_Cond2" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" ("IF _ THEN _ FI" [0,0] 56)
+ (\<open>(0IF _/ THEN _/ ELSE _/ FI)\<close> [0, 0, 0] 61)
+ "_Cond2" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>IF _ THEN _ FI\<close> [0,0] 56)
"_While_inv" :: "'a bexp \<Rightarrow> 'a assn \<Rightarrow> 'a com \<Rightarrow> 'a com"
- ("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61)
+ (\<open>(0WHILE _/ INV _ //DO _ /OD)\<close> [0, 0, 0] 61)
"_While" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"
- ("(0WHILE _ //DO _ /OD)" [0, 0] 61)
+ (\<open>(0WHILE _ //DO _ /OD)\<close> [0, 0] 61)
translations
"IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "CONST Cond \<lbrace>b\<rbrace> c1 c2"
@@ -59,12 +59,12 @@
nonterminal prgs
syntax
- "_PAR" :: "prgs \<Rightarrow> 'a" ("COBEGIN//_//COEND" [57] 56)
- "_prg" :: "['a, 'a] \<Rightarrow> prgs" ("_//_" [60, 90] 57)
- "_prgs" :: "['a, 'a, prgs] \<Rightarrow> prgs" ("_//_//\<parallel>//_" [60,90,57] 57)
+ "_PAR" :: "prgs \<Rightarrow> 'a" (\<open>COBEGIN//_//COEND\<close> [57] 56)
+ "_prg" :: "['a, 'a] \<Rightarrow> prgs" (\<open>_//_\<close> [60, 90] 57)
+ "_prgs" :: "['a, 'a, prgs] \<Rightarrow> prgs" (\<open>_//_//\<parallel>//_\<close> [60,90,57] 57)
"_prg_scheme" :: "['a, 'a, 'a, 'a, 'a] \<Rightarrow> prgs"
- ("SCHEME [_ \<le> _ < _] _// _" [0,0,0,60, 90] 57)
+ (\<open>SCHEME [_ \<le> _ < _] _// _\<close> [0,0,0,60, 90] 57)
translations
"_prg c q" \<rightleftharpoons> "[(CONST Some c, q)]"
--- 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 @@
\<Longrightarrow> interfree (map (\<lambda>k. (c k, Q k)) [a..<b])"
by(force simp add: interfree_def less_diff_conv)
-definition map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \<Rightarrow> bool " ("[\<turnstile>] _" [0] 45) where
+definition map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \<Rightarrow> bool " (\<open>[\<turnstile>] _\<close> [0] 45) where
"[\<turnstile>] Ts == (\<forall>i<length Ts. \<exists>c q. Ts!i=(Some c, q) \<and> \<turnstile> c q)"
lemma MapAnnEmpty: "[\<turnstile>] []"
--- 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 \<times> 'a) \<times> ('a ann_com_op \<times> 'a)) set"
and transition :: "(('a com \<times> 'a) \<times> ('a com \<times> 'a)) set"
and ann_transition' :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
- ("_ -1\<rightarrow> _"[81,81] 100)
+ (\<open>_ -1\<rightarrow> _\<close>[81,81] 100)
and transition' :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
- ("_ -P1\<rightarrow> _"[81,81] 100)
+ (\<open>_ -P1\<rightarrow> _\<close>[81,81] 100)
and transitions :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
- ("_ -P*\<rightarrow> _"[81,81] 100)
+ (\<open>_ -P*\<rightarrow> _\<close>[81,81] 100)
where
"con_0 -1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition"
| "con_0 -P1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition"
@@ -70,17 +70,17 @@
abbreviation
ann_transition_n :: "('a ann_com_op \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a ann_com_op \<times> 'a)
- \<Rightarrow> bool" ("_ -_\<rightarrow> _"[81,81] 100) where
+ \<Rightarrow> bool" (\<open>_ -_\<rightarrow> _\<close>[81,81] 100) where
"con_0 -n\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition ^^ n"
abbreviation
ann_transitions :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
- ("_ -*\<rightarrow> _"[81,81] 100) where
+ (\<open>_ -*\<rightarrow> _\<close>[81,81] 100) where
"con_0 -*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition\<^sup>*"
abbreviation
transition_n :: "('a com \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
- ("_ -P_\<rightarrow> _"[81,81,81] 100) where
+ (\<open>_ -P_\<rightarrow> _\<close>[81,81,81] 100) where
"con_0 -Pn\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition ^^ n"
subsection \<open>Definition of Semantics\<close>
@@ -97,7 +97,7 @@
definition SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set" where
"SEM c S \<equiv> \<Union>(sem c ` S)"
-abbreviation Omega :: "'a com" ("\<Omega>" 63)
+abbreviation Omega :: "'a com" (\<open>\<Omega>\<close> 63)
where "\<Omega> \<equiv> While UNIV UNIV (Basic id)"
primrec fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com" where
@@ -294,10 +294,10 @@
section \<open>Validity of Correctness Formulas\<close>
-definition com_validity :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>= _// _//_)" [90,55,90] 50) where
+definition com_validity :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" (\<open>(3\<parallel>= _// _//_)\<close> [90,55,90] 50) where
"\<parallel>= p c q \<equiv> SEM c p \<subseteq> q"
-definition ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("\<Turnstile> _ _" [60,90] 45) where
+definition ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" (\<open>\<Turnstile> _ _\<close> [60,90] 45) where
"\<Turnstile> c q \<equiv> ann_SEM c (pre c) \<subseteq> q"
end
--- 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 \<Rightarrow> ('a \<Rightarrow> 'b)" ("(\<guillemotleft>_\<guillemotright>)" [0] 1000)
- "_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b" ("\<acute>_" [1000] 1000)
- "_Assert" :: "'a \<Rightarrow> 'a set" ("(\<lbrace>_\<rbrace>)" [0] 1000)
+ "_quote" :: "'b \<Rightarrow> ('a \<Rightarrow> 'b)" (\<open>(\<guillemotleft>_\<guillemotright>)\<close> [0] 1000)
+ "_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b" (\<open>\<acute>_\<close> [1000] 1000)
+ "_Assert" :: "'a \<Rightarrow> 'a set" (\<open>(\<lbrace>_\<rbrace>)\<close> [0] 1000)
translations
"\<lbrace>b\<rbrace>" \<rightharpoonup> "CONST Collect \<guillemotleft>b\<guillemotright>"
--- 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 \<times> 'a) set, ('a \<times> 'a) set, 'a set] \<Rightarrow> bool"
- ("\<turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
+ (\<open>\<turnstile> _ sat [_, _, _, _]\<close> [60,0,0,0,0] 45)
where
Basic: "\<lbrakk> pre \<subseteq> {s. f s \<in> post}; {(s,t). s \<in> pre \<and> (t=f s \<or> t=s)} \<subseteq> guar;
stable pre rely; stable post rely \<rbrakk>
@@ -59,7 +59,7 @@
inductive
par_rghoare :: "('a rgformula) list \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool"
- ("\<turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
+ (\<open>\<turnstile> _ SAT [_, _, _, _]\<close> [60,0,0,0,0] 45)
where
Parallel:
"\<lbrakk> \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j<length xs \<and> j\<noteq>i}. Guar(xs!j)) \<subseteq> Rely(xs!i);
--- 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" (\<open>SKIP\<close>)
where "SKIP \<equiv> Basic id"
-notation Seq ("(_;;/ _)" [60,61] 60)
+notation Seq (\<open>(_;;/ _)\<close> [60,61] 60)
syntax
- "_Assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com" ("(\<acute>_ :=/ _)" [70, 65] 61)
- "_Cond" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com" ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61)
- "_Cond2" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" ("(0IF _ THEN _ FI)" [0,0] 56)
- "_While" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" ("(0WHILE _ /DO _ /OD)" [0, 0] 61)
- "_Await" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" ("(0AWAIT _ /THEN /_ /END)" [0,0] 61)
- "_Atom" :: "'a com \<Rightarrow> 'a com" ("(\<langle>_\<rangle>)" 61)
- "_Wait" :: "'a bexp \<Rightarrow> 'a com" ("(0WAIT _ END)" 61)
+ "_Assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com" (\<open>(\<acute>_ :=/ _)\<close> [70, 65] 61)
+ "_Cond" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0IF _/ THEN _/ ELSE _/FI)\<close> [0, 0, 0] 61)
+ "_Cond2" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0IF _ THEN _ FI)\<close> [0,0] 56)
+ "_While" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0WHILE _ /DO _ /OD)\<close> [0, 0] 61)
+ "_Await" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0AWAIT _ /THEN /_ /END)\<close> [0,0] 61)
+ "_Atom" :: "'a com \<Rightarrow> 'a com" (\<open>(\<langle>_\<rangle>)\<close> 61)
+ "_Wait" :: "'a bexp \<Rightarrow> 'a com" (\<open>(0WAIT _ END)\<close> 61)
translations
"\<acute>x := a" \<rightharpoonup> "CONST Basic \<guillemotleft>\<acute>(_update_name x (\<lambda>_. a))\<guillemotright>"
@@ -30,9 +30,9 @@
nonterminal prgs
syntax
- "_PAR" :: "prgs \<Rightarrow> 'a" ("COBEGIN//_//COEND" 60)
- "_prg" :: "'a \<Rightarrow> prgs" ("_" 57)
- "_prgs" :: "['a, prgs] \<Rightarrow> prgs" ("_//\<parallel>//_" [60,57] 57)
+ "_PAR" :: "prgs \<Rightarrow> 'a" (\<open>COBEGIN//_//COEND\<close> 60)
+ "_prg" :: "'a \<Rightarrow> prgs" (\<open>_\<close> 57)
+ "_prgs" :: "['a, prgs] \<Rightarrow> prgs" (\<open>_//\<parallel>//_\<close> [60,57] 57)
translations
"_prg a" \<rightharpoonup> "[a]"
@@ -40,7 +40,7 @@
"_PAR ps" \<rightharpoonup> "ps"
syntax
- "_prg_scheme" :: "['a, 'a, 'a, 'a] \<Rightarrow> prgs" ("SCHEME [_ \<le> _ < _] _" [0,0,0,60] 57)
+ "_prg_scheme" :: "['a, 'a, 'a, 'a] \<Rightarrow> prgs" (\<open>SCHEME [_ \<le> _ < _] _\<close> [0,0,0,60] 57)
translations
"_prg_scheme j i k c" \<rightleftharpoons> "(CONST map (\<lambda>i. c) [j..<k])"
@@ -48,8 +48,8 @@
text \<open>Translations for variables before and after a transition:\<close>
syntax
- "_before" :: "id \<Rightarrow> 'a" ("\<ordmasculine>_")
- "_after" :: "id \<Rightarrow> 'a" ("\<ordfeminine>_")
+ "_before" :: "id \<Rightarrow> 'a" (\<open>\<ordmasculine>_\<close>)
+ "_after" :: "id \<Rightarrow> 'a" (\<open>\<ordfeminine>_\<close>)
translations
"\<ordmasculine>x" \<rightleftharpoons> "x \<acute>CONST fst"
--- 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 \<times> 'a conf) set"
- and etran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -e\<rightarrow> _" [81,81] 80)
+ and etran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" (\<open>_ -e\<rightarrow> _\<close> [81,81] 80)
where
"P -e\<rightarrow> Q \<equiv> (P,Q) \<in> etran"
| Env: "(P, s) -e\<rightarrow> (P, t)"
@@ -24,8 +24,8 @@
inductive_set
ctran :: "('a conf \<times> 'a conf) set"
- and ctran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -c\<rightarrow> _" [81,81] 80)
- and ctrans :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" ("_ -c*\<rightarrow> _" [81,81] 80)
+ and ctran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" (\<open>_ -c\<rightarrow> _\<close> [81,81] 80)
+ and ctrans :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool" (\<open>_ -c*\<rightarrow> _\<close> [81,81] 80)
where
"P -c\<rightarrow> Q \<equiv> (P,Q) \<in> ctran"
| "P -c*\<rightarrow> Q \<equiv> (P,Q) \<in> ctran\<^sup>*"
@@ -52,14 +52,14 @@
inductive_set
par_etran :: "('a par_conf \<times> 'a par_conf) set"
- and par_etran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pe\<rightarrow> _" [81,81] 80)
+ and par_etran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" (\<open>_ -pe\<rightarrow> _\<close> [81,81] 80)
where
"P -pe\<rightarrow> Q \<equiv> (P,Q) \<in> par_etran"
| ParEnv: "(Ps, s) -pe\<rightarrow> (Ps, t)"
inductive_set
par_ctran :: "('a par_conf \<times> 'a par_conf) set"
- and par_ctran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pc\<rightarrow> _" [81,81] 80)
+ and par_ctran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" (\<open>_ -pc\<rightarrow> _\<close> [81,81] 80)
where
"P -pc\<rightarrow> Q \<equiv> (P,Q) \<in> par_ctran"
| ParComp: "\<lbrakk>i<length Ps; (Ps!i, s) -c\<rightarrow> (r, t)\<rbrakk> \<Longrightarrow> (Ps, s) -pc\<rightarrow> (Ps[i:=r], t)"
@@ -367,7 +367,7 @@
(fst (last c) = None \<longrightarrow> snd (last c) \<in> post)}"
definition com_validity :: "'a com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool"
- ("\<Turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45) where
+ (\<open>\<Turnstile> _ sat [_, _, _, _]\<close> [60,0,0,0,0] 45) where
"\<Turnstile> P sat [pre, rely, guar, post] \<equiv>
\<forall>s. cp (Some P) s \<inter> assum(pre, rely) \<subseteq> comm(guar, post)"
@@ -386,7 +386,7 @@
(All_None (fst (last c)) \<longrightarrow> snd( last c) \<in> post)}"
definition par_com_validity :: "'a par_com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set
-\<Rightarrow> 'a set \<Rightarrow> bool" ("\<Turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45) where
+\<Rightarrow> 'a set \<Rightarrow> bool" (\<open>\<Turnstile> _ SAT [_, _, _, _]\<close> [60,0,0,0,0] 45) where
"\<Turnstile> Ps SAT [pre, rely, guar, post] \<equiv>
\<forall>s. par_cp Ps s \<inter> par_assum(pre, rely) \<subseteq> par_comm(guar, post)"
@@ -409,7 +409,7 @@
(\<forall>l<length clist. l\<noteq>i \<longrightarrow> (clist!l)!j -e\<rightarrow> (clist!l)! Suc j))) \<or>
(c!j -pe\<rightarrow> c!Suc j \<and> (\<forall>i<length clist. (clist!i)!j -e\<rightarrow> (clist!i)! Suc j)))"
-definition conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" ("_ \<propto> _" [65,65] 64) where
+definition conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" (\<open>_ \<propto> _\<close> [65,65] 64) where
"c \<propto> clist \<equiv> (same_length c clist) \<and> (same_state c clist) \<and> (same_program c clist) \<and> (compat_label c clist)"
subsubsection \<open>Some previous lemmas\<close>
--- 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 (\<open>SKIP {_}\<close> 61) |
+ Assign vname aexp 'a (\<open>(_ ::= _/ {_})\<close> [1000, 61, 0] 61) |
+ Seq "('a acom)" "('a acom)" (\<open>_;;//_\<close> [60, 61] 60) |
If bexp 'a "('a acom)" 'a "('a acom)" 'a
- ("(IF _/ THEN ({_}/ _)/ ELSE ({_}/ _)//{_})" [0, 0, 0, 61, 0, 0] 61) |
+ (\<open>(IF _/ THEN ({_}/ _)/ ELSE ({_}/ _)//{_})\<close> [0, 0, 0, 61, 0, 0] 61) |
While 'a bexp 'a "('a acom)" 'a
- ("({_}//WHILE _//DO ({_}//_)//{_})" [0, 0, 0, 61, 0] 61)
+ (\<open>({_}//WHILE _//DO ({_}//_)//{_})\<close> [0, 0, 0, 61, 0] 61)
-notation com.SKIP ("SKIP")
+notation com.SKIP (\<open>SKIP\<close>)
text_raw\<open>\snip{stripdef}{1}{1}{%\<close>
fun strip :: "'a acom \<Rightarrow> com" where
--- 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 \<open>A little syntax magic to write larger states compactly:\<close>
-definition null_state ("<>") where
+definition null_state (\<open><>\<close>) where
"null_state \<equiv> \<lambda>x. 0"
syntax
- "_State" :: "updbinds => 'a" ("<_>")
+ "_State" :: "updbinds => 'a" (\<open><_>\<close>)
translations
"_State ms" == "_Update <> ms"
"_State (_updbinds b bs)" <= "_Update (_State b) bs"
--- 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 \<le> h"
begin
-definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>s") where
+definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" (\<open>m\<^sub>s\<close>) where
"m_s S X = (\<Sum> x \<in> X. m(S x))"
lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> 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 \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
+fun m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" (\<open>m\<^sub>o\<close>) where
"m_o (Some S) X = m_s S X" |
"m_o None X = h * card X + 1"
lemma m_o_h: "finite X \<Longrightarrow> m_o opt X \<le> (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 \<Rightarrow> nat" ("m\<^sub>c") where
+definition m_c :: "'av st option acom \<Rightarrow> nat" (\<open>m\<^sub>c\<close>) where
"m_c C = sum_list (map (\<lambda>a. m_o a (vars C)) (annos C))"
text\<open>Upper complexity bound:\<close>
@@ -359,14 +359,14 @@
the finitely many variables in the program change. That the others do not change
follows because they remain \<^term>\<open>\<top>\<close>.\<close>
-fun top_on_st :: "'av st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>s") where
+fun top_on_st :: "'av st \<Rightarrow> vname set \<Rightarrow> bool" (\<open>top'_on\<^sub>s\<close>) where
"top_on_st S X = (\<forall>x\<in>X. S x = \<top>)"
-fun top_on_opt :: "'av st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>o") where
+fun top_on_opt :: "'av st option \<Rightarrow> vname set \<Rightarrow> bool" (\<open>top'_on\<^sub>o\<close>) 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 \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>c") where
+definition top_on_acom :: "'av st option acom \<Rightarrow> vname set \<Rightarrow> bool" (\<open>top'_on\<^sub>c\<close>) where
"top_on_acom C X = (\<forall>a \<in> set(annos C). top_on_opt a X)"
lemma top_on_top: "top_on_opt \<top> X"
--- 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 \<le> h"
begin
-definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>s") where
+definition m_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" (\<open>m\<^sub>s\<close>) where
"m_s S X = (\<Sum> x \<in> X. m(fun S x))"
lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> 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 \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
+definition m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" (\<open>m\<^sub>o\<close>) where
"m_o opt X = (case opt of None \<Rightarrow> h * card X + 1 | Some S \<Rightarrow> m_s S X)"
lemma m_o_h: "finite X \<Longrightarrow> m_o opt X \<le> (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 \<Rightarrow> nat" ("m\<^sub>c") where
+definition m_c :: "'av st option acom \<Rightarrow> nat" (\<open>m\<^sub>c\<close>) where
"m_c C = sum_list (map (\<lambda>a. m_o a (vars C)) (annos C))"
text\<open>Upper complexity bound:\<close>
@@ -134,14 +134,14 @@
end
-fun top_on_st :: "'a::order_top st \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>s") where
+fun top_on_st :: "'a::order_top st \<Rightarrow> vname set \<Rightarrow> bool" (\<open>top'_on\<^sub>s\<close>) where
"top_on_st S X = (\<forall>x\<in>X. fun S x = \<top>)"
-fun top_on_opt :: "'a::order_top st option \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>o") where
+fun top_on_opt :: "'a::order_top st option \<Rightarrow> vname set \<Rightarrow> bool" (\<open>top'_on\<^sub>o\<close>) 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 \<Rightarrow> vname set \<Rightarrow> bool" ("top'_on\<^sub>c") where
+definition top_on_acom :: "'a::order_top st option acom \<Rightarrow> vname set \<Rightarrow> bool" (\<open>top'_on\<^sub>c\<close>) where
"top_on_acom C X = (\<forall>a \<in> set(annos C). top_on_opt a X)"
lemma top_on_top: "top_on_opt (\<top>::_ st option) X"
--- 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 \<Rightarrow> eint \<Rightarrow> ivl" ("[_, _]") where
+abbreviation ivl_abbr :: "eint \<Rightarrow> eint \<Rightarrow> ivl" (\<open>[_, _]\<close>) where
"[l,h] == abs_ivl(l,h)"
lift_definition \<gamma>_ivl :: "ivl \<Rightarrow> int set" is \<gamma>_rep
--- 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 \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<nabla>" 65)
+fixes widen :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix \<open>\<nabla>\<close> 65)
class narrow =
-fixes narrow :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix "\<triangle>" 65)
+fixes narrow :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infix \<open>\<triangle>\<close> 65)
class wn = widen + narrow + order +
assumes widen1: "x \<le> x \<nabla> y"
@@ -398,7 +398,7 @@
done
-definition n_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" ("n\<^sub>s") where
+definition n_s :: "'av st \<Rightarrow> vname set \<Rightarrow> nat" (\<open>n\<^sub>s\<close>) where
"n\<^sub>s S X = (\<Sum>x\<in>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 \<Rightarrow> vname set \<Rightarrow> nat" ("n\<^sub>o") where
+definition n_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" (\<open>n\<^sub>o\<close>) where
"n\<^sub>o opt X = (case opt of None \<Rightarrow> 0 | Some S \<Rightarrow> n\<^sub>s S X + 1)"
lemma n_o_narrow:
@@ -433,7 +433,7 @@
done
-definition n_c :: "'av st option acom \<Rightarrow> nat" ("n\<^sub>c") where
+definition n_c :: "'av st option acom \<Rightarrow> nat" (\<open>n\<^sub>c\<close>) where
"n\<^sub>c C = sum_list (map (\<lambda>a. n\<^sub>o a (vars C)) (annos C))"
lemma less_annos_iff: "(C1 < C2) = (C1 \<le> C2 \<and>
--- 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\<open>\snip{BigStepdef}{0}{1}{%\<close>
inductive
- big_step :: "com \<times> state \<Rightarrow> state \<Rightarrow> bool" (infix "\<Rightarrow>" 55)
+ big_step :: "com \<times> state \<Rightarrow> state \<Rightarrow> bool" (infix \<open>\<Rightarrow>\<close> 55)
where
Skip: "(SKIP,s) \<Rightarrow> s" |
Assign: "(x ::= a,s) \<Rightarrow> s(x := aval a s)" |
@@ -163,7 +163,7 @@
\<close>
text_raw\<open>\snip{BigStepEquiv}{0}{1}{%\<close>
abbreviation
- equiv_c :: "com \<Rightarrow> com \<Rightarrow> bool" (infix "\<sim>" 50) where
+ equiv_c :: "com \<Rightarrow> com \<Rightarrow> bool" (infix \<open>\<sim>\<close> 50) where
"c \<sim> c' \<equiv> (\<forall>s t. (c,s) \<Rightarrow> t = (c',s) \<Rightarrow> t)"
text_raw\<open>}%endsnip\<close>
--- 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 \<Rightarrow> nat"
-datatype aexp = N nat | Deref aexp ("!") | Plus aexp aexp
+datatype aexp = N nat | Deref aexp (\<open>!\<close>) | Plus aexp aexp
fun aval :: "aexp \<Rightarrow> state \<Rightarrow> nat" where
"aval (N n) s = n" |
@@ -22,14 +22,14 @@
datatype
com = SKIP
- | Assign aexp aexp ("_ ::= _" [61, 61] 61)
+ | Assign aexp aexp (\<open>_ ::= _\<close> [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 (\<open>_;/ _\<close> [60, 61] 60)
+ | If bexp com com (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61)
+ | While bexp com (\<open>(WHILE _/ DO _)\<close> [0, 61] 61)
inductive
- big_step :: "com \<times> state \<times> nat \<Rightarrow> state \<times> nat \<Rightarrow> bool" (infix "\<Rightarrow>" 55)
+ big_step :: "com \<times> state \<times> nat \<Rightarrow> state \<times> nat \<Rightarrow> bool" (infix \<open>\<Rightarrow>\<close> 55)
where
Skip: "(SKIP,sn) \<Rightarrow> sn" |
Assign: "(lhs ::= a,s,n) \<Rightarrow> (s(aval lhs s := aval a s),n)" |
--- 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 "\<squnion>" 65) and
- inf (infixl "\<sqinter>" 70) and
- bot ("\<bottom>") and
- top ("\<top>")
+ sup (infixl \<open>\<squnion>\<close> 65) and
+ inf (infixl \<open>\<sqinter>\<close> 70) and
+ bot (\<open>\<bottom>\<close>) and
+ top (\<open>\<top>\<close>)
context
fixes f :: "vname \<Rightarrow> aexp \<Rightarrow> 'a \<Rightarrow> 'a::sup"
--- 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 (\<open>_ ::= _\<close> [1000, 61] 61)
+ | Seq com com (\<open>_;;/ _\<close> [60, 61] 60)
+ | If bexp com com (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61)
+ | While bexp com (\<open>(WHILE _/ DO _)\<close> [0, 61] 61)
end
--- 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>\<open>i\<close> is an \<^typ>\<open>int\<close>.
\<close>
-fun inth :: "'a list \<Rightarrow> int \<Rightarrow> 'a" (infixl "!!" 100) where
+fun inth :: "'a list \<Rightarrow> int \<Rightarrow> 'a" (infixl \<open>!!\<close> 100) where
"(x # xs) !! i = (if i = 0 then x else xs !! (i - 1))"
text \<open>
@@ -37,7 +37,7 @@
abbreviation (output)
"isize xs == int (length xs)"
-notation isize ("size")
+notation isize (\<open>size\<close>)
subsection "Instructions and Stack Machine"
@@ -66,7 +66,7 @@
definition
exec1 :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool"
- ("(_/ \<turnstile> (_ \<rightarrow>/ _))" [59,0,59] 60)
+ (\<open>(_/ \<turnstile> (_ \<rightarrow>/ _))\<close> [59,0,59] 60)
where
"P \<turnstile> c \<rightarrow> c' =
(\<exists>i s stk. c = (i,s,stk) \<and> c' = iexec(P!!i) (i,s,stk) \<and> 0 \<le> i \<and> i < size P)"
@@ -77,7 +77,7 @@
by (simp add: exec1_def)
abbreviation
- exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" ("(_/ \<turnstile> (_ \<rightarrow>*/ _))" 50)
+ exec :: "instr list \<Rightarrow> config \<Rightarrow> config \<Rightarrow> bool" (\<open>(_/ \<turnstile> (_ \<rightarrow>*/ _))\<close> 50)
where
"exec P \<equiv> star (exec1 P)"
--- 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 \<open>Execution in \<^term>\<open>n\<close> steps for simpler induction\<close>
primrec
exec_n :: "instr list \<Rightarrow> config \<Rightarrow> nat \<Rightarrow> config \<Rightarrow> bool"
- ("_/ \<turnstile> (_ \<rightarrow>^_/ _)" [65,0,1000,55] 55)
+ (\<open>_/ \<turnstile> (_ \<rightarrow>^_/ _)\<close> [65,0,1000,55] 55)
where
"P \<turnstile> c \<rightarrow>^0 c' = (c'=c)" |
"P \<turnstile> c \<rightarrow>^(Suc n) c'' = (\<exists>c'. (P \<turnstile> c \<rightarrow> c') \<and> P \<turnstile> c' \<rightarrow>^n c'')"
--- 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 \<times> state option) \<Rightarrow> state option \<Rightarrow> bool" (infix "\<Rightarrow>" 55)
+ big_step :: "(com \<times> state option) \<Rightarrow> state option \<Rightarrow> bool" (infix \<open>\<Rightarrow>\<close> 55)
where
None: "(c,None) \<Rightarrow> None" |
Skip: "(SKIP,s) \<Rightarrow> s" |
--- 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 \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix "\<rightarrow>" 55)
+ small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix \<open>\<rightarrow>\<close> 55)
where
Assign: "aval a s = Some i \<Longrightarrow> (x ::= a, s) \<rightarrow> (SKIP, s(x := Some i))" |
@@ -21,7 +21,7 @@
lemmas small_step_induct = small_step.induct[split_format(complete)]
-abbreviation small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>*" 55)
+abbreviation small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix \<open>\<rightarrow>*\<close> 55)
where "x \<rightarrow>* y == star small_step x y"
--- 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\<open>Proofs need induction on the length of a small-step reduction sequence.\<close>
fun small_stepsn :: "com * state \<Rightarrow> nat \<Rightarrow> com * state \<Rightarrow> bool"
- ("_ \<rightarrow>'(_') _" [55,0,55] 55) where
+ (\<open>_ \<rightarrow>'(_') _\<close> [55,0,55] 55) where
"(cs \<rightarrow>(0) cs') = (cs' = cs)" |
"cs \<rightarrow>(Suc n) cs'' = (\<exists>cs'. cs \<rightarrow> cs' \<and> cs' \<rightarrow>(n) cs'')"
--- 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 \<Rightarrow> bool"
definition
-hoare_valid :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" ("\<Turnstile> {(1_)}/ (_)/ {(1_)}" 50) where
+hoare_valid :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" (\<open>\<Turnstile> {(1_)}/ (_)/ {(1_)}\<close> 50) where
"\<Turnstile> {P}c{Q} = (\<forall>s t. P s \<and> (c,s) \<Rightarrow> t \<longrightarrow> Q t)"
abbreviation state_subst :: "state \<Rightarrow> aexp \<Rightarrow> vname \<Rightarrow> state"
- ("_[_'/_]" [1000,0,0] 999)
+ (\<open>_[_'/_]\<close> [1000,0,0] 999)
where "s[a/x] == s(x := aval a s)"
inductive
- hoare :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" ("\<turnstile> ({(1_)}/ (_)/ {(1_)})" 50)
+ hoare :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" (\<open>\<turnstile> ({(1_)}/ (_)/ {(1_)})\<close> 50)
where
Skip: "\<turnstile> {P} SKIP {P}" |
--- 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).\<close>
definition hoare_tvalid :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool"
- ("\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where
+ (\<open>\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}\<close> 50) where
"\<Turnstile>\<^sub>t {P}c{Q} \<longleftrightarrow> (\<forall>s. P s \<longrightarrow> (\<exists>t. (c,s) \<Rightarrow> t \<and> Q t))"
text\<open>Provability of Hoare triples in the proof system for total
@@ -22,7 +22,7 @@
\<^term>\<open>While\<close>-rule.\<close>
inductive
- hoaret :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" ("\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50)
+ hoaret :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" (\<open>\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})\<close> 50)
where
Skip: "\<turnstile>\<^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:\<close>
-definition wpt :: "com \<Rightarrow> assn \<Rightarrow> assn" ("wp\<^sub>t") where
+definition wpt :: "com \<Rightarrow> assn \<Rightarrow> assn" (\<open>wp\<^sub>t\<close>) where
"wp\<^sub>t c Q = (\<lambda>s. \<exists>t. (c,s) \<Rightarrow> t \<and> Q t)"
lemma [simp]: "wp\<^sub>t SKIP Q = Q"
--- 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.\<close>
definition hoare_tvalid :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool"
- ("\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where
+ (\<open>\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}\<close> 50) where
"\<Turnstile>\<^sub>t {P}c{Q} \<longleftrightarrow> (\<forall>s. P s \<longrightarrow> (\<exists>t. (c,s) \<Rightarrow> t \<and> Q t))"
inductive
- hoaret :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" ("\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50)
+ hoaret :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" (\<open>\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})\<close> 50)
where
Skip: "\<turnstile>\<^sub>t {P} SKIP {P}" |
@@ -68,7 +68,7 @@
qed fastforce+
-definition wpt :: "com \<Rightarrow> assn \<Rightarrow> assn" ("wp\<^sub>t") where
+definition wpt :: "com \<Rightarrow> assn \<Rightarrow> assn" (\<open>wp\<^sub>t\<close>) where
"wp\<^sub>t c Q = (\<lambda>s. \<exists>t. (c,s) \<Rightarrow> t \<and> Q t)"
lemma [simp]: "wp\<^sub>t SKIP Q = Q"
--- 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 \<Rightarrow> nat) \<Rightarrow> state \<Rightarrow> bool"
definition hoare_tvalid :: "assn2 \<Rightarrow> com \<Rightarrow> assn2 \<Rightarrow> bool"
- ("\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where
+ (\<open>\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}\<close> 50) where
"\<Turnstile>\<^sub>t {P}c{Q} \<longleftrightarrow> (\<forall>l s. P l s \<longrightarrow> (\<exists>t. (c,s) \<Rightarrow> t \<and> Q l t))"
inductive
- hoaret :: "assn2 \<Rightarrow> com \<Rightarrow> assn2 \<Rightarrow> bool" ("\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50)
+ hoaret :: "assn2 \<Rightarrow> com \<Rightarrow> assn2 \<Rightarrow> bool" (\<open>\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})\<close> 50)
where
Skip: "\<turnstile>\<^sub>t {P} SKIP {P}" |
@@ -75,7 +75,7 @@
qed fastforce+
-definition wpt :: "com \<Rightarrow> assn2 \<Rightarrow> assn2" ("wp\<^sub>t") where
+definition wpt :: "com \<Rightarrow> assn2 \<Rightarrow> assn2" (\<open>wp\<^sub>t\<close>) where
"wp\<^sub>t c Q = (\<lambda>l s. \<exists>t. (c,s) \<Rightarrow> t \<and> Q l t)"
lemma [simp]: "wp\<^sub>t SKIP Q = Q"
--- 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 \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c"
- ("_/'((2_,_ :=/ _)')" [1000,0,0,0] 900)
+ (\<open>_/'((2_,_ :=/ _)')\<close> [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 ("_\<bullet>/_" [63,1000] 63) |
- Vassign string exp ("(_ ::=/ _)" [1000,61] 62) |
- Fassign exp string exp ("(_\<bullet>_ ::=/ _)" [63,0,62] 62) |
- Mcall exp string exp ("(_\<bullet>/_<_>)" [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 (\<open>_\<bullet>/_\<close> [63,1000] 63) |
+ Vassign string exp (\<open>(_ ::=/ _)\<close> [1000,61] 62) |
+ Fassign exp string exp (\<open>(_\<bullet>_ ::=/ _)\<close> [63,0,62] 62) |
+ Mcall exp string exp (\<open>(_\<bullet>/_<_>)\<close> [63,0,0] 63) |
+ Seq exp exp (\<open>_;/ _\<close> [61,60] 60) |
+ If bexp exp exp (\<open>IF _/ THEN (2_)/ ELSE (2_)\<close> [0,0,61] 61)
and bexp = B bool | Not bexp | And bexp bexp | Eq exp exp
type_synonym menv = "string \<Rightarrow> exp"
@@ -31,9 +31,9 @@
inductive
big_step :: "menv \<Rightarrow> exp \<times> config \<Rightarrow> ref \<times> config \<Rightarrow> bool"
- ("(_ \<turnstile>/ (_/ \<Rightarrow> _))" [60,0,60] 55) and
+ (\<open>(_ \<turnstile>/ (_/ \<Rightarrow> _))\<close> [60,0,60] 55) and
bval :: "menv \<Rightarrow> bexp \<times> config \<Rightarrow> bool \<times> config \<Rightarrow> bool"
- ("_ \<turnstile> _ \<rightarrow> _" [60,0,60] 55)
+ (\<open>_ \<turnstile> _ \<rightarrow> _\<close> [60,0,60] 55)
where
Null:
"me \<turnstile> (Null,c) \<Rightarrow> (null,c)" |
--- 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 \<Rightarrow> ty"
inductive atyping :: "tyenv \<Rightarrow> aexp \<Rightarrow> ty \<Rightarrow> bool"
- ("(1_/ \<turnstile>p/ (_ :/ _))" [50,0,50] 50)
+ (\<open>(1_/ \<turnstile>p/ (_ :/ _))\<close> [50,0,50] 50)
where
"\<Gamma> \<turnstile>p Ic i : Ity" |
"\<Gamma> \<turnstile>p Rc r : Rty" |
"\<Gamma> \<turnstile>p V x : \<Gamma> x" |
"\<Gamma> \<turnstile>p a1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile>p a2 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile>p Plus a1 a2 : \<tau>"
-inductive btyping :: "tyenv \<Rightarrow> bexp \<Rightarrow> bool" (infix "\<turnstile>p" 50)
+inductive btyping :: "tyenv \<Rightarrow> bexp \<Rightarrow> bool" (infix \<open>\<turnstile>p\<close> 50)
where
"\<Gamma> \<turnstile>p Bc v" |
"\<Gamma> \<turnstile>p b \<Longrightarrow> \<Gamma> \<turnstile>p Not b" |
"\<Gamma> \<turnstile>p b1 \<Longrightarrow> \<Gamma> \<turnstile>p b2 \<Longrightarrow> \<Gamma> \<turnstile>p And b1 b2" |
"\<Gamma> \<turnstile>p a1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile>p a2 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile>p Less a1 a2"
-inductive ctyping :: "tyenv \<Rightarrow> com \<Rightarrow> bool" (infix "\<turnstile>p" 50) where
+inductive ctyping :: "tyenv \<Rightarrow> com \<Rightarrow> bool" (infix \<open>\<turnstile>p\<close> 50) where
"\<Gamma> \<turnstile>p SKIP" |
"\<Gamma> \<turnstile>p a : \<Gamma>(x) \<Longrightarrow> \<Gamma> \<turnstile>p x ::= a" |
"\<Gamma> \<turnstile>p c1 \<Longrightarrow> \<Gamma> \<turnstile>p c2 \<Longrightarrow> \<Gamma> \<turnstile>p c1;;c2" |
@@ -34,7 +34,7 @@
"type (Iv i) = Ity" |
"type (Rv r) = Rty"
-definition styping :: "tyenv \<Rightarrow> state \<Rightarrow> bool" (infix "\<turnstile>p" 50)
+definition styping :: "tyenv \<Rightarrow> state \<Rightarrow> bool" (infix \<open>\<turnstile>p\<close> 50)
where "\<Gamma> \<turnstile>p s \<longleftrightarrow> (\<forall>x. type (s x) = \<Gamma> x)"
fun tsubst :: "(nat \<Rightarrow> ty) \<Rightarrow> ty \<Rightarrow> ty" where
--- 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 (\<open>_ ::= _\<close> [1000, 61] 61)
+ | Seq com com (\<open>_;;/ _\<close> [60, 61] 60)
+ | If bexp com com (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61)
+ | While bexp com (\<open>(WHILE _/ DO _)\<close> [0, 61] 61)
+ | Var vname com (\<open>(1{VAR _;/ _})\<close>)
+ | Proc pname com com (\<open>(1{PROC _ = _;/ _})\<close>)
| CALL pname
definition "test_com =
--- 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 \<Rightarrow> com"
inductive
- big_step :: "penv \<Rightarrow> com \<times> state \<Rightarrow> state \<Rightarrow> bool" ("_ \<turnstile> _ \<Rightarrow> _" [60,0,60] 55)
+ big_step :: "penv \<Rightarrow> com \<times> state \<Rightarrow> state \<Rightarrow> bool" (\<open>_ \<turnstile> _ \<Rightarrow> _\<close> [60,0,60] 55)
where
Skip: "pe \<turnstile> (SKIP,s) \<Rightarrow> s" |
Assign: "pe \<turnstile> (x ::= a,s) \<Rightarrow> s(x := aval a s)" |
--- 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 \<times> com) list"
inductive
- big_step :: "penv \<Rightarrow> com \<times> state \<Rightarrow> state \<Rightarrow> bool" ("_ \<turnstile> _ \<Rightarrow> _" [60,0,60] 55)
+ big_step :: "penv \<Rightarrow> com \<times> state \<Rightarrow> state \<Rightarrow> bool" (\<open>_ \<turnstile> _ \<Rightarrow> _\<close> [60,0,60] 55)
where
Skip: "pe \<turnstile> (SKIP,s) \<Rightarrow> s" |
Assign: "pe \<turnstile> (x ::= a,s) \<Rightarrow> s(x := aval a s)" |
--- 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 \<times> venv \<times> nat \<Rightarrow> com \<times> store \<Rightarrow> store \<Rightarrow> bool"
- ("_ \<turnstile> _ \<Rightarrow> _" [60,0,60] 55)
+ (\<open>_ \<turnstile> _ \<Rightarrow> _\<close> [60,0,60] 55)
where
Skip: "e \<turnstile> (SKIP,s) \<Rightarrow> s" |
Assign: "(pe,ve,f) \<turnstile> (x ::= a,s) \<Rightarrow> s(ve x := aval a (s o ve))" |
--- 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 \<Rightarrow> state \<Rightarrow> level \<Rightarrow> bool"
- ("(_ = _ '(\<le> _'))" [51,51,0] 50) where
+ (\<open>(_ = _ '(\<le> _'))\<close> [51,51,0] 50) where
"s = s' (\<le> l) == (\<forall> x. sec x \<le> l \<longrightarrow> s x = s' x)"
abbreviation eq_less :: "state \<Rightarrow> state \<Rightarrow> level \<Rightarrow> bool"
- ("(_ = _ '(< _'))" [51,51,0] 50) where
+ (\<open>(_ = _ '(< _'))\<close> [51,51,0] 50) where
"s = s' (< l) == (\<forall> x. sec x < l \<longrightarrow> s x = s' x)"
lemma aval_eq_if_eq_le:
--- 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 \<Rightarrow> com \<Rightarrow> bool" ("(_/ \<turnstile> _)" [0,0] 50) where
+inductive sec_type :: "nat \<Rightarrow> com \<Rightarrow> bool" (\<open>(_/ \<turnstile> _)\<close> [0,0] 50) where
Skip:
"l \<turnstile> SKIP" |
Assign:
@@ -183,7 +183,7 @@
computation by an antimonotonicity rule. We introduce the standard system now
and show the equivalence with our formulation.\<close>
-inductive sec_type' :: "nat \<Rightarrow> com \<Rightarrow> bool" ("(_/ \<turnstile>'' _)" [0,0] 50) where
+inductive sec_type' :: "nat \<Rightarrow> com \<Rightarrow> bool" (\<open>(_/ \<turnstile>'' _)\<close> [0,0] 50) where
Skip':
"l \<turnstile>' SKIP" |
Assign':
@@ -217,7 +217,7 @@
subsubsection "A Bottom-Up Typing System"
-inductive sec_type2 :: "com \<Rightarrow> level \<Rightarrow> bool" ("(\<turnstile> _ : _)" [0,0] 50) where
+inductive sec_type2 :: "com \<Rightarrow> level \<Rightarrow> bool" (\<open>(\<turnstile> _ : _)\<close> [0,0] 50) where
Skip2:
"\<turnstile> SKIP : l" |
Assign2:
--- 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 \<Rightarrow> com \<Rightarrow> bool" ("(_/ \<turnstile> _)" [0,0] 50) where
+inductive sec_type :: "nat \<Rightarrow> com \<Rightarrow> bool" (\<open>(_/ \<turnstile> _)\<close> [0,0] 50) where
Skip:
"l \<turnstile> SKIP" |
Assign:
@@ -172,7 +172,7 @@
computation by an antimonotonicity rule. We introduce the standard system now
and show the equivalence with our formulation.\<close>
-inductive sec_type' :: "nat \<Rightarrow> com \<Rightarrow> bool" ("(_/ \<turnstile>'' _)" [0,0] 50) where
+inductive sec_type' :: "nat \<Rightarrow> com \<Rightarrow> bool" (\<open>(_/ \<turnstile>'' _)\<close> [0,0] 50) where
Skip':
"l \<turnstile>' SKIP" |
Assign':
--- 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 \<Rightarrow> bool"
definition
- equiv_up_to :: "assn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> bool" ("_ \<Turnstile> _ \<sim> _" [50,0,10] 50)
+ equiv_up_to :: "assn \<Rightarrow> com \<Rightarrow> com \<Rightarrow> bool" (\<open>_ \<Turnstile> _ \<sim> _\<close> [50,0,10] 50)
where
"(P \<Turnstile> c \<sim> c') = (\<forall>s s'. P s \<longrightarrow> (c,s) \<Rightarrow> s' \<longleftrightarrow> (c',s) \<Rightarrow> s')"
definition
- bequiv_up_to :: "assn \<Rightarrow> bexp \<Rightarrow> bexp \<Rightarrow> bool" ("_ \<Turnstile> _ <\<sim>> _" [50,0,10] 50)
+ bequiv_up_to :: "assn \<Rightarrow> bexp \<Rightarrow> bexp \<Rightarrow> bool" (\<open>_ \<Turnstile> _ <\<sim>> _\<close> [50,0,10] 50)
where
"(P \<Turnstile> b <\<sim>> b') = (\<forall>s. P s \<longrightarrow> bval b s = bval b' s)"
--- 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 \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>" 55)
+ small_step :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix \<open>\<rightarrow>\<close> 55)
where
Assign: "(x ::= a, s) \<rightarrow> (SKIP, s(x := aval a s))" |
@@ -20,7 +20,7 @@
abbreviation
- small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>*" 55)
+ small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix \<open>\<rightarrow>*\<close> 55)
where "x \<rightarrow>* y == star small_step x y"
subsection\<open>Executability\<close>
--- 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 (\<open>_ ::= _\<close> [1000, 61] 61)
+ | Seq com com (\<open>_;; _\<close> [60, 61] 60)
+ | If bexp com com (\<open>IF _ THEN _ ELSE _\<close> [0, 0, 61] 61)
+ | While bexp com (\<open>WHILE _ DO _\<close> [0, 61] 61)
subsection "Small-Step Semantics of Commands"
inductive
- small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix "\<rightarrow>" 55)
+ small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix \<open>\<rightarrow>\<close> 55)
where
Assign: "taval a s v \<Longrightarrow> (x ::= a, s) \<rightarrow> (SKIP, s(x := v))" |
@@ -76,7 +76,7 @@
type_synonym tyenv = "vname \<Rightarrow> ty"
inductive atyping :: "tyenv \<Rightarrow> aexp \<Rightarrow> ty \<Rightarrow> bool"
- ("(1_/ \<turnstile>/ (_ :/ _))" [50,0,50] 50)
+ (\<open>(1_/ \<turnstile>/ (_ :/ _))\<close> [50,0,50] 50)
where
Ic_ty: "\<Gamma> \<turnstile> Ic i : Ity" |
Rc_ty: "\<Gamma> \<turnstile> 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.\<close>
-inductive btyping :: "tyenv \<Rightarrow> bexp \<Rightarrow> bool" (infix "\<turnstile>" 50)
+inductive btyping :: "tyenv \<Rightarrow> bexp \<Rightarrow> bool" (infix \<open>\<turnstile>\<close> 50)
where
B_ty: "\<Gamma> \<turnstile> Bc v" |
Not_ty: "\<Gamma> \<turnstile> b \<Longrightarrow> \<Gamma> \<turnstile> Not b" |
@@ -102,7 +102,7 @@
declare btyping.intros [intro!]
inductive_cases [elim!]: "\<Gamma> \<turnstile> Not b" "\<Gamma> \<turnstile> And b1 b2" "\<Gamma> \<turnstile> Less a1 a2"
-inductive ctyping :: "tyenv \<Rightarrow> com \<Rightarrow> bool" (infix "\<turnstile>" 50) where
+inductive ctyping :: "tyenv \<Rightarrow> com \<Rightarrow> bool" (infix \<open>\<turnstile>\<close> 50) where
Skip_ty: "\<Gamma> \<turnstile> SKIP" |
Assign_ty: "\<Gamma> \<turnstile> a : \<Gamma>(x) \<Longrightarrow> \<Gamma> \<turnstile> x ::= a" |
Seq_ty: "\<Gamma> \<turnstile> c1 \<Longrightarrow> \<Gamma> \<turnstile> c2 \<Longrightarrow> \<Gamma> \<turnstile> c1;;c2" |
@@ -127,7 +127,7 @@
lemma type_eq_Rty[simp]: "type v = Rty \<longleftrightarrow> (\<exists>r. v = Rv r)"
by (cases v) simp_all
-definition styping :: "tyenv \<Rightarrow> state \<Rightarrow> bool" (infix "\<turnstile>" 50)
+definition styping :: "tyenv \<Rightarrow> state \<Rightarrow> bool" (infix \<open>\<turnstile>\<close> 50)
where "\<Gamma> \<turnstile> s \<longleftrightarrow> (\<forall>x. type (s x) = \<Gamma> x)"
lemma apreservation:
@@ -204,7 +204,7 @@
"(c,s) \<rightarrow> (c',s') \<Longrightarrow> \<Gamma> \<turnstile> c \<Longrightarrow> \<Gamma> \<turnstile> c'"
by (induct rule: small_step_induct) (auto simp: ctyping.intros)
-abbreviation small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>*" 55)
+abbreviation small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix \<open>\<rightarrow>*\<close> 55)
where "x \<rightarrow>* y == star small_step x y"
theorem type_sound:
--- 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\<open>Commands where loops are annotated with invariants.\<close>
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 (\<open>SKIP\<close>) |
+ Aassign vname aexp (\<open>(_ ::= _)\<close> [1000, 61] 61) |
+ Aseq acom acom (\<open>_;;/ _\<close> [60, 61] 60) |
+ Aif bexp acom acom (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61) |
+ Awhile assn bexp acom (\<open>({_}/ WHILE _/ DO _)\<close> [0, 0, 61] 61)
-notation com.SKIP ("SKIP")
+notation com.SKIP (\<open>SKIP\<close>)
text\<open>Strip annotations:\<close>
--- 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.\<close>
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 (\<open>SKIP\<close>) |
+ Aassign vname aexp (\<open>(_ ::= _)\<close> [1000, 61] 61) |
+ Aseq acom acom (\<open>_;;/ _\<close> [60, 61] 60) |
+ Aif bexp acom acom (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61) |
Awhile "nat \<Rightarrow> assn" bexp acom
- ("({_}/ WHILE _/ DO _)" [0, 0, 61] 61)
+ (\<open>({_}/ WHILE _/ DO _)\<close> [0, 0, 61] 61)
-notation com.SKIP ("SKIP")
+notation com.SKIP (\<open>SKIP\<close>)
text\<open>Strip annotations:\<close>
--- 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.\<close>
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 (\<open>SKIP\<close>) |
+ Aassign vname aexp (\<open>(_ ::= _)\<close> [1000, 61] 61) |
+ Aseq acom acom (\<open>_;;/ _\<close> [60, 61] 60) |
+ Aif bexp acom acom (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61) |
Awhile assn2 lvname bexp acom
- ("({_'/_}/ WHILE _/ DO _)" [0, 0, 0, 61] 61)
+ (\<open>({_'/_}/ WHILE _/ DO _)\<close> [0, 0, 0, 61] 61)
-notation com.SKIP ("SKIP")
+notation com.SKIP (\<open>SKIP\<close>)
text\<open>Strip annotations:\<close>
--- 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 \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> bool"
- ("(_ =/ _/ on _)" [50,0,50] 50) where
+ (\<open>(_ =/ _/ on _)\<close> [50,0,50] 50) where
"f = g on X == \<forall> x \<in> X. f x = g x"
lemma aval_eq_if_eq_on_vars[simp]:
--- 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 (\<open>_:==_\<close> [65, 65 ] 60)
+ | Local loc aexp com (\<open>LOCAL _:=_ IN _\<close> [65, 0, 61] 60)
+ | Semi com com (\<open>_;; _\<close> [59, 60 ] 59)
+ | Cond bexp com com (\<open>IF _ THEN _ ELSE _\<close> [65, 60, 61] 60)
+ | While bexp com (\<open>WHILE _ DO _\<close> [65, 61] 60)
| BODY pname
- | Call vname pname aexp ("_:=CALL _'(_')" [65, 65, 0] 60)
+ | Call vname pname aexp (\<open>_:=CALL _'(_')\<close> [65, 65, 0] 60)
consts bodies :: "(pname * com) list"(* finitely many procedure definitions *)
definition
--- 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" (\<open>Z=Arg+_\<close> [50]50) where
"Z=Arg+n = (%Z s. Z = s<Arg>+n)"
definition
--- 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 = (\<exists>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 \<open>&>\<close> 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" (\<open>{(1_)}./ (_)/ .{(1_)}\<close> [3,60,3] 58)
definition
- triple_valid :: "nat => 'a triple => bool" ( "|=_:_" [0 , 58] 57) where
+ triple_valid :: "nat => 'a triple => bool" ( \<open>|=_:_\<close> [0 , 58] 57) where
"|=n:t = (case t of {P}.c.{Q} =>
\<forall>Z s. P Z s \<longrightarrow> (\<forall>s'. <c,s> -n-> s' \<longrightarrow> Q Z s'))"
abbreviation
- triples_valid :: "nat => 'a triple set => bool" ("||=_:_" [0 , 58] 57) where
+ triples_valid :: "nat => 'a triple set => bool" (\<open>||=_:_\<close> [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" (\<open>_||=_\<close> [58, 58] 57) where
"G||=ts = (\<forall>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" (\<open>_|=_\<close> [58, 58] 57) where
"G |=t == G||={t}"
(* Most General Triples *)
definition
- MGT :: "com => state triple" ("{=}._.{->}" [60] 58) where
+ MGT :: "com => state triple" (\<open>{=}._.{->}\<close> [60] 58) where
"{=}.c.{->} = {%Z s0. Z = s0}. c .{%Z s1. <c,Z> -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" (\<open>_||-_\<close> [58, 58] 57) and
+ hoare_deriv :: "'a triple set => 'a triple => bool" (\<open>_|-_\<close> [58, 58] 57)
where
"G |-t == G||-{t}"
--- 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" (\<open>_/[_/::=/_]\<close> [900,0,0] 900)
abbreviation
- loc :: "state => locals" ("_<_>" [75,0] 75) where
+ loc :: "state => locals" (\<open>_<_>\<close> [75,0] 75) where
"s<X> == getlocs s X"
inductive
- evalc :: "[com,state, state] => bool" ("<_,_>/ -c-> _" [0,0, 51] 51)
+ evalc :: "[com,state, state] => bool" (\<open><_,_>/ -c-> _\<close> [0,0, 51] 51)
where
Skip: "<SKIP,s> -c-> s"
@@ -52,7 +52,7 @@
[X::=s1<Res>]"
inductive
- evaln :: "[com,state,nat,state] => bool" ("<_,_>/ -_-> _" [0,0,0,51] 51)
+ evaln :: "[com,state,nat,state] => bool" (\<open><_,_>/ -_-> _\<close> [0,0,0,51] 51)
where
Skip: "<SKIP,s> -n-> s"
--- 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) \<union> outputs(a2)),
(internals(a1) \<union> internals(a2))))"
-definition par :: "[('a,'s)ioa, ('a,'t)ioa] \<Rightarrow> ('a,'s*'t)ioa" (infixr "||" 10)
+definition par :: "[('a,'s)ioa, ('a,'t)ioa] \<Rightarrow> ('a,'s*'t)ioa" (infixr \<open>||\<close> 10)
where "(ioa1 || ioa2) \<equiv>
(asig_comp (asig_of ioa1) (asig_of ioa2),
{pr. fst(pr) \<in> starts_of(ioa1) \<and> snd(pr) \<in> starts_of(ioa2)},
--- 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 \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> heap \<Rightarrow> heap" where
"update a i x h = set a ((get h a)[i:=x]) h"
-definition noteq :: "'a::heap array \<Rightarrow> 'b::heap array \<Rightarrow> bool" (infix "=!!=" 70) where
+definition noteq :: "'a::heap array \<Rightarrow> 'b::heap array \<Rightarrow> bool" (infix \<open>=!!=\<close> 70) where
"r =!!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_array r \<noteq> addr_of_array s"
--- 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" (\<open>_::_\<close> [4, 0] 3)
+ "_constrain" :: "prop' => type => prop'" (\<open>_::_\<close> [4, 0] 3)
syntax (output)
- "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3)
- "_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3)
+ "_constrain" :: "logic => type => logic" (\<open>_ :: _\<close> [4, 0] 3)
+ "_constrain" :: "prop' => type => prop'" (\<open>_ :: _\<close> [4, 0] 3)
(*>*)
text \<open>
--- 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\<lparr>lim := l + 1\<rparr>)))"
-definition noteq :: "'a::heap ref \<Rightarrow> 'b::heap ref \<Rightarrow> bool" (infix "=!=" 70) where
+definition noteq :: "'a::heap ref \<Rightarrow> 'b::heap ref \<Rightarrow> bool" (infix \<open>=!=\<close> 70) where
"r =!= s \<longleftrightarrow> TYPEREP('a) \<noteq> TYPEREP('b) \<or> addr_of_ref r \<noteq> addr_of_ref s"
@@ -41,10 +41,10 @@
definition ref :: "'a::heap \<Rightarrow> 'a ref Heap" where
[code del]: "ref v = Heap_Monad.heap (alloc v)"
-definition lookup :: "'a::heap ref \<Rightarrow> 'a Heap" ("!_" 61) where
+definition lookup :: "'a::heap ref \<Rightarrow> 'a Heap" (\<open>!_\<close> 61) where
[code del]: "lookup r = Heap_Monad.tap (\<lambda>h. get h r)"
-definition update :: "'a ref \<Rightarrow> 'a::heap \<Rightarrow> unit Heap" ("_ := _" 62) where
+definition update :: "'a ref \<Rightarrow> 'a::heap \<Rightarrow> unit Heap" (\<open>_ := _\<close> 62) where
[code del]: "update r v = Heap_Monad.heap (\<lambda>h. ((), set r v h))"
definition change :: "('a::heap \<Rightarrow> 'a) \<Rightarrow> 'a ref \<Rightarrow> 'a Heap" where
--- 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 (\<open>VALOF _ RESULTIS _\<close> 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 \<open>:=\<close> 60)
+ | Semi com com (\<open>_;;_\<close> [60, 60] 60)
+ | Cond exp com com (\<open>IF _ THEN _ ELSE _\<close> 60)
+ | While exp com (\<open>WHILE _ DO _\<close> 60)
subsection \<open>Commands\<close>
@@ -30,7 +30,7 @@
text\<open>Execution of commands\<close>
abbreviation (input)
- generic_rel ("_/ -|[_]-> _" [50,0,50] 50) where
+ generic_rel (\<open>_/ -|[_]-> _\<close> [50,0,50] 50) where
"esig -|[eval]-> ns == (esig,ns) \<in> eval"
text\<open>Command execution. Natural numbers represent Booleans: 0=True, 1=False\<close>
@@ -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)
+ (\<open>_/ -[_]-> _\<close> [50,0,50] 50)
for eval :: "((exp*state) * (nat*state)) set"
where
"csig -[eval]-> s == (csig,s) \<in> 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 \<open>-|->\<close> 50)
where
"esig -|-> ns == (esig, ns) \<in> eval"
--- 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 "\<bullet>" 90)
+ | Ap comb comb (infixl \<open>\<bullet>\<close> 90)
text \<open>
Inductive definition of contractions, \<open>\<rightarrow>\<^sup>1\<close> and
(multi-step) reductions, \<open>\<rightarrow>\<close>.
\<close>
-inductive contract1 :: "[comb,comb] \<Rightarrow> bool" (infixl "\<rightarrow>\<^sup>1" 50)
+inductive contract1 :: "[comb,comb] \<Rightarrow> bool" (infixl \<open>\<rightarrow>\<^sup>1\<close> 50)
where
K: "K\<bullet>x\<bullet>y \<rightarrow>\<^sup>1 x"
| S: "S\<bullet>x\<bullet>y\<bullet>z \<rightarrow>\<^sup>1 (x\<bullet>z)\<bullet>(y\<bullet>z)"
@@ -35,7 +35,7 @@
| Ap2: "x \<rightarrow>\<^sup>1 y \<Longrightarrow> z\<bullet>x \<rightarrow>\<^sup>1 z\<bullet>y"
abbreviation
- contract :: "[comb,comb] \<Rightarrow> bool" (infixl "\<rightarrow>" 50) where
+ contract :: "[comb,comb] \<Rightarrow> bool" (infixl \<open>\<rightarrow>\<close> 50) where
"contract \<equiv> contract1\<^sup>*\<^sup>*"
text \<open>
@@ -43,7 +43,7 @@
(multi-step) parallel reductions, \<open>\<Rrightarrow>\<close>.
\<close>
-inductive parcontract1 :: "[comb,comb] \<Rightarrow> bool" (infixl "\<Rrightarrow>\<^sup>1" 50)
+inductive parcontract1 :: "[comb,comb] \<Rightarrow> bool" (infixl \<open>\<Rrightarrow>\<^sup>1\<close> 50)
where
refl: "x \<Rrightarrow>\<^sup>1 x"
| K: "K\<bullet>x\<bullet>y \<Rrightarrow>\<^sup>1 x"
@@ -51,7 +51,7 @@
| Ap: "\<lbrakk>x \<Rrightarrow>\<^sup>1 y; z \<Rrightarrow>\<^sup>1 w\<rbrakk> \<Longrightarrow> x\<bullet>z \<Rrightarrow>\<^sup>1 y\<bullet>w"
abbreviation
- parcontract :: "[comb,comb] \<Rightarrow> bool" (infixl "\<Rrightarrow>" 50) where
+ parcontract :: "[comb,comb] \<Rightarrow> bool" (infixl \<open>\<Rrightarrow>\<close> 50) where
"parcontract \<equiv> parcontract1\<^sup>*\<^sup>*"
text \<open>
--- 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 \<Rightarrow> (ordinal \<Rightarrow> ordinal)) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"
where "OpLim F a = Limit (\<lambda>n. F n a)"
-definition OpItw :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" ("\<Squnion>")
+definition OpItw :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" (\<open>\<Squnion>\<close>)
where "\<Squnion>f = OpLim (iter f)"
primrec cantor :: "ordinal \<Rightarrow> ordinal \<Rightarrow> ordinal"
@@ -40,7 +40,7 @@
| "cantor a (Succ b) = \<Squnion>(\<lambda>x. cantor x b) a"
| "cantor a (Limit f) = Limit (\<lambda>n. cantor a (f n))"
-primrec Nabla :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" ("\<nabla>")
+primrec Nabla :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" (\<open>\<nabla>\<close>)
where
"\<nabla>f Zero = f Zero"
| "\<nabla>f (Succ a) = f (Succ (\<nabla>f a))"
--- 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 "\<rightharpoonup>" 90)
+ | var 'a (\<open>#_\<close> [1000])
+ | imp "'a pl" "'a pl" (infixr \<open>\<rightharpoonup>\<close> 90)
subsection \<open>The proof system\<close>
-inductive thms :: "['a pl set, 'a pl] \<Rightarrow> bool" (infixl "\<turnstile>" 50)
+inductive thms :: "['a pl set, 'a pl] \<Rightarrow> bool" (infixl \<open>\<turnstile>\<close> 50)
for H :: "'a pl set"
where
H: "p \<in> H \<Longrightarrow> H \<turnstile> p"
@@ -42,7 +42,7 @@
subsubsection \<open>Semantics of propositional logic.\<close>
-primrec eval :: "['a set, 'a pl] => bool" ("_[[_]]" [100,0] 100)
+primrec eval :: "['a set, 'a pl] => bool" (\<open>_[[_]]\<close> [100,0] 100)
where
"tt[[false]] = False"
| "tt[[#v]] = (v \<in> tt)"
@@ -67,7 +67,7 @@
is \<open>p\<close>.
\<close>
-definition sat :: "['a pl set, 'a pl] => bool" (infixl "\<Turnstile>" 50)
+definition sat :: "['a pl set, 'a pl] => bool" (infixl \<open>\<Turnstile>\<close> 50)
where "H \<Turnstile> p = (\<forall>tt. (\<forall>q\<in>H. tt[[q]]) \<longrightarrow> tt[[p]])"
--- 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 "\<sim>" 50)
+ and msg_rel :: "[freemsg, freemsg] => bool" (infixl \<open>\<sim>\<close> 50)
where
"X \<sim> Y == (X,Y) \<in> msgrel"
| CD: "CRYPT K (DECRYPT K X) \<sim> X"
--- 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.\<close>
inductive_set
exprel :: "(freeExp * freeExp) set"
- and exp_rel :: "[freeExp, freeExp] => bool" (infixl "\<sim>" 50)
+ and exp_rel :: "[freeExp, freeExp] => bool" (infixl \<open>\<sim>\<close> 50)
where
"X \<sim> Y \<equiv> (X,Y) \<in> exprel"
| ASSOC: "PLUS X (PLUS Y Z) \<sim> PLUS (PLUS X Y) Z"
--- 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 (\<open>[]\<close>) and
+ Cons (infixr \<open>#\<close> 65)
definition
- Nil :: "'a list" ("[]") where
+ Nil :: "'a list" (\<open>[]\<close>) where
"Nil = Abs_List(NIL)"
definition
- "Cons" :: "['a, 'a list] => 'a list" (infixr "#" 65) where
+ "Cons" :: "['a, 'a list] => 'a list" (infixr \<open>#\<close> 65) where
"x#xs = Abs_List(CONS (Leaf x)(Rep_List xs))"
definition
--- 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 \<open>hypothetical group axiomatization\<close>
context
- fixes prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<odot>" 70)
+ fixes prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<odot>\<close> 70)
and one :: "'a"
and inverse :: "'a \<Rightarrow> 'a"
assumes assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)"
--- 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 \<open>hypothetical group axiomatization\<close>
- fix prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<odot>" 70)
+ fix prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<odot>\<close> 70)
and one :: "'a"
and inverse :: "'a \<Rightarrow> 'a"
assume assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)"
--- 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 \<Rightarrow> 'a"
- | Seq "'a com" "'a com" ("(_;/ _)" [60, 61] 60)
+ | Seq "'a com" "'a com" (\<open>(_;/ _)\<close> [60, 61] 60)
| Cond "'a bexp" "'a com" "'a com"
| While "'a bexp" "'a assn" "'a var" "'a com"
-abbreviation Skip ("SKIP")
+abbreviation Skip (\<open>SKIP\<close>)
where "SKIP \<equiv> Basic id"
type_synonym 'a sem = "'a \<Rightarrow> 'a \<Rightarrow> bool"
@@ -46,7 +46,7 @@
| "Sem (Cond b c1 c2) s s' \<longleftrightarrow> (if s \<in> b then Sem c1 s s' else Sem c2 s s')"
| "Sem (While b x y c) s s' \<longleftrightarrow> (\<exists>n. iter n b (Sem c) s s')"
-definition Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool" ("(3\<turnstile> _/ (2_)/ _)" [100, 55, 100] 50)
+definition Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool" (\<open>(3\<turnstile> _/ (2_)/ _)\<close> [100, 55, 100] 50)
where "\<turnstile> P c Q \<longleftrightarrow> (\<forall>s s'. Sem c s s' \<longrightarrow> s \<in> P \<longrightarrow> s' \<in> Q)"
lemma ValidI [intro?]: "(\<And>s s'. Sem c s s' \<Longrightarrow> s \<in> P \<Longrightarrow> s' \<in> Q) \<Longrightarrow> \<turnstile> P c Q"
@@ -193,15 +193,15 @@
syntax
"_quote" :: "'b \<Rightarrow> ('a \<Rightarrow> 'b)"
- "_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b" ("\<acute>_" [1000] 1000)
- "_Subst" :: "'a bexp \<Rightarrow> 'b \<Rightarrow> idt \<Rightarrow> 'a bexp" ("_[_'/\<acute>_]" [1000] 999)
- "_Assert" :: "'a \<Rightarrow> 'a set" ("(\<lbrace>_\<rbrace>)" [0] 1000)
- "_Assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com" ("(\<acute>_ :=/ _)" [70, 65] 61)
+ "_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b" (\<open>\<acute>_\<close> [1000] 1000)
+ "_Subst" :: "'a bexp \<Rightarrow> 'b \<Rightarrow> idt \<Rightarrow> 'a bexp" (\<open>_[_'/\<acute>_]\<close> [1000] 999)
+ "_Assert" :: "'a \<Rightarrow> 'a set" (\<open>(\<lbrace>_\<rbrace>)\<close> [0] 1000)
+ "_Assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com" (\<open>(\<acute>_ :=/ _)\<close> [70, 65] 61)
"_Cond" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"
- ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61)
+ (\<open>(0IF _/ THEN _/ ELSE _/ FI)\<close> [0, 0, 0] 61)
"_While_inv" :: "'a bexp \<Rightarrow> 'a assn \<Rightarrow> 'a com \<Rightarrow> 'a com"
- ("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61)
- "_While" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" ("(0WHILE _ //DO _ /OD)" [0, 0] 61)
+ (\<open>(0WHILE _/ INV _ //DO _ /OD)\<close> [0, 0, 0] 61)
+ "_While" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com" (\<open>(0WHILE _ //DO _ /OD)\<close> [0, 0] 61)
translations
"\<lbrace>b\<rbrace>" \<rightharpoonup> "CONST Collect (_quote b)"
--- 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 @@
\<close>
definition
- Meet :: "'a::complete_lattice set \<Rightarrow> 'a" ("\<Sqinter>_" [90] 90) where
+ Meet :: "'a::complete_lattice set \<Rightarrow> 'a" (\<open>\<Sqinter>_\<close> [90] 90) where
"\<Sqinter>A = (THE inf. is_Inf A inf)"
definition
- Join :: "'a::complete_lattice set \<Rightarrow> 'a" ("\<Squnion>_" [90] 90) where
+ Join :: "'a::complete_lattice set \<Rightarrow> 'a" (\<open>\<Squnion>_\<close> [90] 90) where
"\<Squnion>A = (THE sup. is_Sup A sup)"
text \<open>
@@ -184,11 +184,11 @@
\<close>
definition
- bottom :: "'a::complete_lattice" ("\<bottom>") where
+ bottom :: "'a::complete_lattice" (\<open>\<bottom>\<close>) where
"\<bottom> = \<Sqinter>UNIV"
definition
- top :: "'a::complete_lattice" ("\<top>") where
+ top :: "'a::complete_lattice" (\<open>\<top>\<close>) where
"\<top> = \<Squnion>UNIV"
lemma bottom_least [intro?]: "\<bottom> \<sqsubseteq> x"
--- 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 @@
\<close>
definition
- meet :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70) where
+ meet :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<sqinter>\<close> 70) where
"x \<sqinter> y = (THE inf. is_inf x y inf)"
definition
- join :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65) where
+ join :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<squnion>\<close> 65) where
"x \<squnion> y = (THE sup. is_sup x y sup)"
text \<open>
--- 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 @@
\<close>
class leq =
- fixes leq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sqsubseteq>" 50)
+ fixes leq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>\<sqsubseteq>\<close> 50)
class quasi_order = leq +
assumes leq_refl [intro?]: "x \<sqsubseteq> x"
--- 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 \<open>Cardinalities of types\<close>
-syntax "_type_card" :: "type => nat" ("(1CARD/(1'(_')))")
+syntax "_type_card" :: "type => nat" (\<open>(1CARD/(1'(_')))\<close>)
syntax_consts "_type_card" == card
--- 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 = (\<lambda>x y. P x \<and> P y) \<sqinter> R"
lemma combine_per_simp [simp]:
- "combine_per P R x y \<longleftrightarrow> P x \<and> P y \<and> x \<approx> y" for R (infixl "\<approx>" 50)
+ "combine_per P R x y \<longleftrightarrow> P x \<and> P y \<and> x \<approx> y" for R (infixl \<open>\<approx>\<close> 50)
by (simp add: combine_per_def)
lemma combine_per_top [simp]: "combine_per \<top> R = R"
@@ -27,7 +27,7 @@
lemma transp_combine_per: "transp R \<Longrightarrow> transp (combine_per P R)"
by (auto simp add: transp_def trans_def combine_per_def)
-lemma combine_perI: "P x \<Longrightarrow> P y \<Longrightarrow> x \<approx> y \<Longrightarrow> combine_per P R x y" for R (infixl "\<approx>" 50)
+lemma combine_perI: "P x \<Longrightarrow> P y \<Longrightarrow> x \<approx> y \<Longrightarrow> combine_per P R x y" for R (infixl \<open>\<approx>\<close> 50)
by (simp add: combine_per_def)
lemma symp_combine_per_symp: "symp R \<Longrightarrow> 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 "\<approx>" 50)
+ fixes R (infixl \<open>\<approx>\<close> 50)
assumes "\<exists>x. P x" and "equivp R"
shows "part_equivp (combine_per P R)"
proof -
--- 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 "\<sqsubseteq>" 60)
+ fixes le_b (infix \<open>\<sqsubseteq>\<close> 60)
assumes chain: "Complete_Partial_Order.chain (fun_ord (\<le>)) Y"
and mono: "\<And>f. f \<in> Y \<Longrightarrow> monotone le_b (\<le>) f"
shows "monotone (\<sqsubseteq>) (\<le>) (fun_lub Sup Y)"
@@ -96,7 +96,7 @@
qed
context
- fixes le_b (infix "\<sqsubseteq>" 60) and Y f
+ fixes le_b (infix \<open>\<sqsubseteq>\<close> 60) and Y f
assumes chain: "Complete_Partial_Order.chain le_b Y"
and mono1: "\<And>y. y \<in> Y \<Longrightarrow> monotone le_b (\<le>) (\<lambda>x. f x y)"
and mono2: "\<And>x a b. \<lbrakk> x \<in> Y; a \<sqsubseteq> b; a \<in> Y; b \<in> Y \<rbrakk> \<Longrightarrow> f x a \<le> f x b"
@@ -163,7 +163,7 @@
end
lemma Sup_image_mono_le:
- fixes le_b (infix "\<sqsubseteq>" 60) and Sup_b ("\<Or>")
+ fixes le_b (infix \<open>\<sqsubseteq>\<close> 60) and Sup_b (\<open>\<Or>\<close>)
assumes ccpo: "class.ccpo Sup_b (\<sqsubseteq>) lt_b"
assumes chain: "Complete_Partial_Order.chain (\<sqsubseteq>) Y"
and mono: "\<And>x y. \<lbrakk> x \<sqsubseteq> y; x \<in> Y \<rbrakk> \<Longrightarrow> f x \<le> f y"
@@ -181,7 +181,7 @@
qed
lemma swap_Sup:
- fixes le_b (infix "\<sqsubseteq>" 60)
+ fixes le_b (infix \<open>\<sqsubseteq>\<close> 60)
assumes Y: "Complete_Partial_Order.chain (\<sqsubseteq>) Y"
and Z: "Complete_Partial_Order.chain (fun_ord (\<le>)) Z"
and mono: "\<And>f. f \<in> Z \<Longrightarrow> monotone (\<sqsubseteq>) (\<le>) f"
@@ -289,7 +289,7 @@
qed(blast intro: ccpo_Sup_least)
qed(rule chain_iterates[OF f])
-context fixes ordb :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix "\<sqsubseteq>" 60) begin
+context fixes ordb :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix \<open>\<sqsubseteq>\<close> 60) begin
lemma iterates_mono:
assumes f: "f \<in> ccpo.iterates (fun_lub Sup) (fun_ord (\<le>)) F"
@@ -569,8 +569,8 @@
by(rule mcont2mcont'[OF _ mcont_const])
context
- fixes ord :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix "\<sqsubseteq>" 60)
- and lub :: "'b set \<Rightarrow> 'b" ("\<Or>")
+ fixes ord :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix \<open>\<sqsubseteq>\<close> 60)
+ and lub :: "'b set \<Rightarrow> 'b" (\<open>\<Or>\<close>)
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 ("\<Or>") and ord (infix "\<sqsubseteq>" 60)
+ fixes bot and lub (\<open>\<Or>\<close>) and ord (infix \<open>\<sqsubseteq>\<close> 60)
assumes ccpo: "class.ccpo lub (\<sqsubseteq>) lt"
and mono: "\<And>x y. \<lbrakk> x \<le> y; \<not> x \<le> bound \<rbrakk> \<Longrightarrow> f x \<sqsubseteq> f y"
and cont: "\<And>Y. \<lbrakk> Complete_Partial_Order.chain (\<le>) Y; Y \<noteq> {}; \<And>x. x \<in> Y \<Longrightarrow> \<not> x \<le> bound \<rbrakk> \<Longrightarrow> f (\<Squnion>Y) = \<Or>(f ` Y)"
@@ -891,7 +891,7 @@
end
lemma admissible_leI:
- fixes ord (infix "\<sqsubseteq>" 60) and lub ("\<Or>")
+ fixes ord (infix \<open>\<sqsubseteq>\<close> 60) and lub (\<open>\<Or>\<close>)
assumes "class.ccpo lub (\<sqsubseteq>) (mk_less (\<sqsubseteq>))"
and "mcont luba orda lub (\<sqsubseteq>) (\<lambda>x. f x)"
and "mcont luba orda lub (\<sqsubseteq>) (\<lambda>x. g x)"
--- 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 @@
\<close>
no_syntax
- "_constify" :: "id => ident" ("_")
- "_constify" :: "longid => ident" ("_")
+ "_constify" :: "id => ident" (\<open>_\<close>)
+ "_constify" :: "longid => ident" (\<open>_\<close>)
- "_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\<lparr>_\<rparr>)")
- "_record_type_scheme" :: "field_types => type => type" ("(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)")
+ "_field_type" :: "ident => type => field_type" (\<open>(2_ ::/ _)\<close>)
+ "" :: "field_type => field_types" (\<open>_\<close>)
+ "_field_types" :: "field_type => field_types => field_types" (\<open>_,/ _\<close>)
+ "_record_type" :: "field_types => type" (\<open>(3\<lparr>_\<rparr>)\<close>)
+ "_record_type_scheme" :: "field_types => type => type" (\<open>(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)\<close>)
- "_field" :: "ident => 'a => field" ("(2_ =/ _)")
- "" :: "field => fields" ("_")
- "_fields" :: "field => fields => fields" ("_,/ _")
- "_record" :: "fields => 'a" ("(3\<lparr>_\<rparr>)")
- "_record_scheme" :: "fields => 'a => 'a" ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
+ "_field" :: "ident => 'a => field" (\<open>(2_ =/ _)\<close>)
+ "" :: "field => fields" (\<open>_\<close>)
+ "_fields" :: "field => fields => fields" (\<open>_,/ _\<close>)
+ "_record" :: "fields => 'a" (\<open>(3\<lparr>_\<rparr>)\<close>)
+ "_record_scheme" :: "fields => 'a => 'a" (\<open>(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)\<close>)
- "_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\<lparr>_\<rparr>)" [900, 0] 900)
+ "_field_update" :: "ident => 'a => field_update" (\<open>(2_ :=/ _)\<close>)
+ "" :: "field_update => field_updates" (\<open>_\<close>)
+ "_field_updates" :: "field_update => field_updates => field_updates" (\<open>_,/ _\<close>)
+ "_record_update" :: "'a => field_updates => 'b" (\<open>_/(3\<lparr>_\<rparr>)\<close> [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" (\<open>(3'(| _ |'))\<close>)
+ "_record_type_scheme" :: "field_types => type => type" (\<open>(3'(| _,/ (2... ::/ _) |'))\<close>)
+ "_record" :: "fields => 'a" (\<open>(3'(| _ |'))\<close>)
+ "_record_scheme" :: "fields => 'a => 'a" (\<open>(3'(| _,/ (2... =/ _) |'))\<close>)
+ "_record_update" :: "'a => field_updates => 'b" (\<open>_/(3'(| _ |'))\<close> [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" (\<open>_\<close>)
+ "_constify" :: "longid => ident" (\<open>_\<close>)
- "_datatype_field" :: "ident => 'a => field" ("(2_ =/ _)")
- "" :: "field => fields" ("_")
- "_datatype_fields" :: "field => fields => fields" ("_,/ _")
- "_datatype_record" :: "fields => 'a" ("(3\<lparr>_\<rparr>)")
- "_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\<lparr>_\<rparr>)" [900, 0] 900)
+ "_datatype_field" :: "ident => 'a => field" (\<open>(2_ =/ _)\<close>)
+ "" :: "field => fields" (\<open>_\<close>)
+ "_datatype_fields" :: "field => fields => fields" (\<open>_,/ _\<close>)
+ "_datatype_record" :: "fields => 'a" (\<open>(3\<lparr>_\<rparr>)\<close>)
+ "_datatype_field_update" :: "ident => 'a => field_update" (\<open>(2_ :=/ _)\<close>)
+ "" :: "field_update => field_updates" (\<open>_\<close>)
+ "_datatype_field_updates" :: "field_update => field_updates => field_updates" (\<open>_,/ _\<close>)
+ "_datatype_record_update" :: "'a => field_updates => 'b" (\<open>_/(3\<lparr>_\<rparr>)\<close> [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" (\<open>(3'(| _ |'))\<close>)
+ "_datatype_record_scheme" :: "fields => 'a => 'a" (\<open>(3'(| _,/ (2... =/ _) |'))\<close>)
+ "_datatype_record_update" :: "'a => field_updates => 'b" (\<open>_/(3'(| _ |'))\<close> [900, 0] 900)
named_theorems datatype_record_update
--- 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\<open>Eqpoll\<close>
-definition eqpoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl "\<approx>" 50)
+definition eqpoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl \<open>\<approx>\<close> 50)
where "eqpoll A B \<equiv> \<exists>f. bij_betw f A B"
-definition lepoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl "\<lesssim>" 50)
+definition lepoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl \<open>\<lesssim>\<close> 50)
where "lepoll A B \<equiv> \<exists>f. inj_on f A \<and> f ` A \<subseteq> B"
definition lesspoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl \<open>\<prec>\<close> 50)
--- 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 ("\<infinity>") | Minf ("-\<infinity>")
+datatype 'a extended = Fin 'a | Pinf (\<open>\<infinity>\<close>) | Minf (\<open>-\<infinity>\<close>)
instantiation extended :: (order)order
--- 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" ("\<infinity>")
+ fixes infinity :: "'a" (\<open>\<infinity>\<close>)
context
fixes f :: "nat \<Rightarrow> 'a::{canonically_ordered_monoid_add, linorder_topology, complete_linorder}"
--- 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 "{||} \<equiv> bot"
-abbreviation fsubset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50) where "xs |\<subseteq>| ys \<equiv> xs \<le> ys"
-abbreviation fsubset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50) where "xs |\<subset>| ys \<equiv> xs < ys"
-abbreviation funion :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|\<union>|" 65) where "xs |\<union>| ys \<equiv> sup xs ys"
-abbreviation finter :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|\<inter>|" 65) where "xs |\<inter>| ys \<equiv> inf xs ys"
-abbreviation fminus :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl "|-|" 65) where "xs |-| ys \<equiv> minus xs ys"
+abbreviation fempty :: "'a fset" (\<open>{||}\<close>) where "{||} \<equiv> bot"
+abbreviation fsubset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<subseteq>|\<close> 50) where "xs |\<subseteq>| ys \<equiv> xs \<le> ys"
+abbreviation fsubset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<subset>|\<close> 50) where "xs |\<subset>| ys \<equiv> xs < ys"
+abbreviation funion :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl \<open>|\<union>|\<close> 65) where "xs |\<union>| ys \<equiv> sup xs ys"
+abbreviation finter :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl \<open>|\<inter>|\<close> 65) where "xs |\<inter>| ys \<equiv> inf xs ys"
+abbreviation fminus :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" (infixl \<open>|-|\<close> 65) where "xs |-| ys \<equiv> minus xs ys"
instantiation fset :: (equal) equal
begin
@@ -154,7 +154,7 @@
end
abbreviation fUNIV :: "'a::finite fset" where "fUNIV \<equiv> top"
-abbreviation fuminus :: "'a::finite fset \<Rightarrow> 'a fset" ("|-| _" [81] 80) where "|-| x \<equiv> uminus x"
+abbreviation fuminus :: "'a::finite fset \<Rightarrow> 'a fset" (\<open>|-| _\<close> [81] 80) where "|-| x \<equiv> uminus x"
declare top_fset.rep_eq[simp]
@@ -166,19 +166,19 @@
nonterminal fset_args
syntax
- "" :: "'a \<Rightarrow> fset_args" ("_")
- "_fset_args" :: "'a \<Rightarrow> fset_args \<Rightarrow> fset_args" ("_,/ _")
- "_fset" :: "fset_args => 'a fset" ("{|(_)|}")
+ "" :: "'a \<Rightarrow> fset_args" (\<open>_\<close>)
+ "_fset_args" :: "'a \<Rightarrow> fset_args \<Rightarrow> fset_args" (\<open>_,/ _\<close>)
+ "_fset" :: "fset_args => 'a fset" (\<open>{|(_)|}\<close>)
syntax_consts
"_fset_args" "_fset" == finsert
translations
"{|x, xs|}" == "CONST finsert x {|xs|}"
"{|x|}" == "CONST finsert x {||}"
-abbreviation fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<in>|" 50) where
+abbreviation fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<in>|\<close> 50) where
"x |\<in>| X \<equiv> x \<in> fset X"
-abbreviation not_fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50) where
+abbreviation not_fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<notin>|\<close> 50) where
"x |\<notin>| X \<equiv> x \<notin> fset X"
context
@@ -197,12 +197,12 @@
end
syntax (input)
- "_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3! (_/|:|_)./ _)" [0, 0, 10] 10)
- "_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3? (_/|:|_)./ _)" [0, 0, 10] 10)
+ "_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3! (_/|:|_)./ _)\<close> [0, 0, 10] 10)
+ "_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3? (_/|:|_)./ _)\<close> [0, 0, 10] 10)
syntax
- "_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<forall>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
- "_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" ("(3\<exists>(_/|\<in>|_)./ _)" [0, 0, 10] 10)
+ "_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<forall>(_/|\<in>|_)./ _)\<close> [0, 0, 10] 10)
+ "_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<exists>(_/|\<in>|_)./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_fBall" \<rightleftharpoons> FSet.Ball and
@@ -243,7 +243,7 @@
lift_definition fcard :: "'a fset \<Rightarrow> nat" is card parametric card_transfer .
-lift_definition fimage :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" (infixr "|`|" 90) is image
+lift_definition fimage :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" (infixr \<open>|`|\<close> 90) is image
parametric image_transfer by simp
lift_definition fthe_elem :: "'a fset \<Rightarrow> 'a" is the_elem .
@@ -1905,7 +1905,7 @@
text \<open>Setup adapted from sets.\<close>
-notation Quickcheck_Exhaustive.orelse (infixr "orelse" 55)
+notation Quickcheck_Exhaustive.orelse (infixr \<open>orelse\<close> 55)
context
includes term_syntax
@@ -1939,7 +1939,7 @@
end
-no_notation Quickcheck_Exhaustive.orelse (infixr "orelse" 55)
+no_notation Quickcheck_Exhaustive.orelse (infixr \<open>orelse\<close> 55)
instantiation fset :: (random) random
begin
--- 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 \<Rightarrow> ('a, 'b) fmap \<Rightarrow> ('a, 'b) fmap" (infixl "++\<^sub>f" 100)
+lift_definition fmadd :: "('a, 'b) fmap \<Rightarrow> ('a, 'b) fmap \<Rightarrow> ('a, 'b) fmap" (infixl \<open>++\<^sub>f\<close> 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 \<Rightarrow> ('a, 'b) fmap \<Rightarrow> bool" (infix "\<subseteq>\<^sub>f" 50)
+lift_definition fmsubset :: "('a, 'b) fmap \<Rightarrow> ('a, 'b) fmap \<Rightarrow> bool" (infix \<open>\<subseteq>\<^sub>f\<close> 50)
is map_le
.
@@ -801,7 +801,7 @@
obtains x where "fmlookup m x = Some y" "x |\<in>| A"
using assms by (auto simp: fmlookup_image_iff)
-lift_definition fmcomp :: "('b, 'c) fmap \<Rightarrow> ('a, 'b) fmap \<Rightarrow> ('a, 'c) fmap" (infixl "\<circ>\<^sub>f" 55)
+lift_definition fmcomp :: "('b, 'c) fmap \<Rightarrow> ('a, 'b) fmap \<Rightarrow> ('a, 'c) fmap" (infixl \<open>\<circ>\<^sub>f\<close> 55)
is map_comp
parametric map_comp_transfer
by (rule dom_comp_finite)
--- 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 \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> 'b"
where "restrict f A = (\<lambda>x. if x \<in> A then f x else undefined)"
-abbreviation funcset :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" (infixr "\<rightarrow>" 60)
+abbreviation funcset :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" (infixr \<open>\<rightarrow>\<close> 60)
where "A \<rightarrow> B \<equiv> Pi A (\<lambda>_. B)"
syntax
- "_Pi" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" ("(3\<Pi> _\<in>_./ _)" 10)
- "_lam" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)" ("(3\<lambda>_\<in>_./ _)" [0,0,3] 3)
+ "_Pi" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" (\<open>(3\<Pi> _\<in>_./ _)\<close> 10)
+ "_lam" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b)" (\<open>(3\<lambda>_\<in>_./ _)\<close> [0,0,3] 3)
syntax_consts
"_Pi" \<rightleftharpoons> Pi and
"_lam" \<rightleftharpoons> restrict
@@ -345,13 +345,13 @@
abbreviation "Pi\<^sub>E A B \<equiv> PiE A B"
syntax
- "_PiE" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" ("(3\<Pi>\<^sub>E _\<in>_./ _)" 10)
+ "_PiE" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" (\<open>(3\<Pi>\<^sub>E _\<in>_./ _)\<close> 10)
syntax_consts
"_PiE" \<rightleftharpoons> Pi\<^sub>E
translations
"\<Pi>\<^sub>E x\<in>A. B" \<rightleftharpoons> "CONST Pi\<^sub>E A (\<lambda>x. B)"
-abbreviation extensional_funcset :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" (infixr "\<rightarrow>\<^sub>E" 60)
+abbreviation extensional_funcset :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a \<Rightarrow> 'b) set" (infixr \<open>\<rightarrow>\<^sub>E\<close> 60)
where "A \<rightarrow>\<^sub>E B \<equiv> (\<Pi>\<^sub>E i\<in>A. B)"
lemma extensional_funcset_def: "extensional_funcset S T = (S \<rightarrow> T) \<inter> extensional S"
--- 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 \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3SUM _. _)" [0, 10] 10)
+ "_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" (\<open>(3SUM _. _)\<close> [0, 10] 10)
syntax
- "_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3\<Sum>_. _)" [0, 10] 10)
+ "_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" (\<open>(3\<Sum>_. _)\<close> [0, 10] 10)
syntax_consts
"_Sum_any" \<rightleftharpoons> Sum_any
translations
@@ -254,9 +254,9 @@
end
syntax (ASCII)
- "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" ("(3PROD _. _)" [0, 10] 10)
+ "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" (\<open>(3PROD _. _)\<close> [0, 10] 10)
syntax
- "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" ("(3\<Prod>_. _)" [0, 10] 10)
+ "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" (\<open>(3\<Prod>_. _)\<close> [0, 10] 10)
syntax_consts
"_Prod_any" == Prod_any
translations
--- 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 \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a iarray" where
[simp]: "of_fun f n = IArray (map f [0..<n])"
-qualified definition sub :: "'a iarray \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!!" 100) where
+qualified definition sub :: "'a iarray \<Rightarrow> nat \<Rightarrow> 'a" (infixl \<open>!!\<close> 100) where
[simp]: "as !! n = IArray.list_of as ! n"
qualified definition length :: "'a iarray \<Rightarrow> nat" where
--- 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 \<open>Membership\<close>
-abbreviation (in preorder) in_interval ("(_/ \<in>\<^sub>i _)" [51, 51] 50)
+abbreviation (in preorder) in_interval (\<open>(_/ \<in>\<^sub>i _)\<close> [51, 51] 50)
where "in_interval x X \<equiv> x \<in> set_of X"
lemma in_interval_to_interval[intro!]: "a \<in>\<^sub>i interval_of a"
--- 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 \<in>\<^sub>i X" if "lower X \<le> x" "x \<le> upper X"
using that by (auto simp: set_of_eq)
-abbreviation in_real_interval ("(_/ \<in>\<^sub>r _)" [51, 51] 50) where
+abbreviation in_real_interval (\<open>(_/ \<in>\<^sub>r _)\<close> [51, 51] 50) where
"x \<in>\<^sub>r X \<equiv> x \<in>\<^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 \<Rightarrow> 'a interval list \<Rightarrow> bool"
- (infix "(all'_in\<^sub>i)" 50)
+ (infix \<open>(all'_in\<^sub>i)\<close> 50)
where "x all_in\<^sub>i I = (length x = length I \<and> (\<forall>i < length I. x ! i \<in>\<^sub>i I ! i))"
definition all_in :: "real list \<Rightarrow> float interval list \<Rightarrow> bool"
- (infix "(all'_in)" 50)
+ (infix \<open>(all'_in)\<close> 50)
where "x all_in I = (length x = length I \<and> (\<forall>i < length I. x ! i \<in>\<^sub>r I ! i))"
definition all_subset :: "'a::order interval list \<Rightarrow> 'a interval list \<Rightarrow> bool"
- (infix "(all'_subset)" 50)
+ (infix \<open>(all'_subset)\<close> 50)
where "I all_subset J = (length I = length J \<and> (\<forall>i < length I. set_of (I!i) \<subseteq> set_of (J!i)))"
lemmas [simp] = all_in_def all_subset_def
--- 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 \<Rightarrow> 'a" where
"mbox0 x = x"
-notation (latex) mbox ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>" [999] 999)
+notation (latex) mbox (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close> [999] 999)
-notation (latex) mbox0 ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>" [0] 999)
+notation (latex) mbox0 (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close> [0] 999)
(* LOGIC *)
notation (latex output)
- If ("(\<^latex>\<open>\\textsf{\<close>if\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\\textsf{\<close>then\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\\textsf{\<close>else\<^latex>\<open>}\<close> (_))" 10)
+ If (\<open>(\<^latex>\<open>\textsf{\<close>if\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\textsf{\<close>then\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\textsf{\<close>else\<^latex>\<open>}\<close> (_))\<close> 10)
syntax (latex output)
"_Let" :: "[letbinds, 'a] => 'a"
- ("(\<^latex>\<open>\\textsf{\<close>let\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\\textsf{\<close>in\<^latex>\<open>}\<close> (_))" 10)
+ (\<open>(\<^latex>\<open>\textsf{\<close>let\<^latex>\<open>}\<close> (_)/ \<^latex>\<open>\textsf{\<close>in\<^latex>\<open>}\<close> (_))\<close> 10)
"_case_syntax":: "['a, cases_syn] => 'b"
- ("(\<^latex>\<open>\\textsf{\<close>case\<^latex>\<open>}\<close> _ \<^latex>\<open>\\textsf{\<close>of\<^latex>\<open>}\<close>/ _)" 10)
+ (\<open>(\<^latex>\<open>\textsf{\<close>case\<^latex>\<open>}\<close> _ \<^latex>\<open>\textsf{\<close>of\<^latex>\<open>}\<close>/ _)\<close> 10)
(* SETS *)
(* empty set *)
notation (latex)
- "Set.empty" ("\<emptyset>")
+ "Set.empty" (\<open>\<emptyset>\<close>)
(* 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{_ \<in> _ | _})")
+ "_Collect" :: "pttrn => bool => 'a set" (\<open>(1{_ | _})\<close>)
+ "_CollectIn" :: "pttrn => 'a set => bool => 'a set" (\<open>(1{_ \<in> _ | _})\<close>)
translations
"_Collect p P" <= "{p. P}"
"_Collect p P" <= "{p|xs. P}"
@@ -57,56 +57,56 @@
(* card *)
notation (latex output)
- card ("|_|")
+ card (\<open>|_|\<close>)
(* LISTS *)
(* Cons *)
notation (latex)
- Cons ("_ \<cdot>/ _" [66,65] 65)
+ Cons (\<open>_ \<cdot>/ _\<close> [66,65] 65)
(* length *)
notation (latex output)
- length ("|_|")
+ length (\<open>|_|\<close>)
(* nth *)
notation (latex output)
- nth ("_\<^latex>\<open>\\ensuremath{_{[\\mathit{\<close>_\<^latex>\<open>}]}}\<close>" [1000,0] 1000)
+ nth (\<open>_\<^latex>\<open>\ensuremath{_{[\mathit{\<close>_\<^latex>\<open>}]}}\<close>\<close> [1000,0] 1000)
(* DUMMY *)
-consts DUMMY :: 'a ("\<^latex>\<open>\\_\<close>")
+consts DUMMY :: 'a (\<open>\<^latex>\<open>\_\<close>\<close>)
(* THEOREMS *)
notation (Rule output)
- Pure.imp ("\<^latex>\<open>\\mbox{}\\inferrule{\\mbox{\<close>_\<^latex>\<open>}}\<close>\<^latex>\<open>{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+ Pure.imp (\<open>\<^latex>\<open>\mbox{}\inferrule{\mbox{\<close>_\<^latex>\<open>}}\<close>\<^latex>\<open>{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
syntax (Rule output)
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
- ("\<^latex>\<open>\\mbox{}\\inferrule{\<close>_\<^latex>\<open>}\<close>\<^latex>\<open>{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+ (\<open>\<^latex>\<open>\mbox{}\inferrule{\<close>_\<^latex>\<open>}\<close>\<^latex>\<open>{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
"_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms"
- ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\\\\\<close>/ _")
+ (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\\\<close>/ _\<close>)
- "_asm" :: "prop \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>")
+ "_asm" :: "prop \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close>)
notation (Axiom output)
- "Trueprop" ("\<^latex>\<open>\\mbox{}\\inferrule{\\mbox{}}{\\mbox{\<close>_\<^latex>\<open>}}\<close>")
+ "Trueprop" (\<open>\<^latex>\<open>\mbox{}\inferrule{\mbox{}}{\mbox{\<close>_\<^latex>\<open>}}\<close>\<close>)
notation (IfThen output)
- Pure.imp ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _/ \<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
+ Pure.imp (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _/ \<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
syntax (IfThen output)
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
- ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _ /\<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
- "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close> /\<^latex>\<open>{\\normalsize \\,\<close>and\<^latex>\<open>\\,}\<close>/ _")
- "_asm" :: "prop \<Rightarrow> asms" ("\<^latex>\<open>\\mbox{\<close>_\<^latex>\<open>}\<close>")
+ (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _ /\<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
+ "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close> /\<^latex>\<open>{\normalsize \,\<close>and\<^latex>\<open>\,}\<close>/ _\<close>)
+ "_asm" :: "prop \<Rightarrow> asms" (\<open>\<^latex>\<open>\mbox{\<close>_\<^latex>\<open>}\<close>\<close>)
notation (IfThenNoBox output)
- Pure.imp ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _/ \<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
+ Pure.imp (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _/ \<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
syntax (IfThenNoBox output)
"_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
- ("\<^latex>\<open>{\\normalsize{}\<close>If\<^latex>\<open>\\,}\<close> _ /\<^latex>\<open>{\\normalsize \\,\<close>then\<^latex>\<open>\\,}\<close>/ _.")
- "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^latex>\<open>{\\normalsize \\,\<close>and\<^latex>\<open>\\,}\<close>/ _")
- "_asm" :: "prop \<Rightarrow> asms" ("_")
+ (\<open>\<^latex>\<open>{\normalsize{}\<close>If\<^latex>\<open>\,}\<close> _ /\<^latex>\<open>{\normalsize \,\<close>then\<^latex>\<open>\,}\<close>/ _.\<close>)
+ "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" (\<open>_ /\<^latex>\<open>{\normalsize \,\<close>and\<^latex>\<open>\,}\<close>/ _\<close>)
+ "_asm" :: "prop \<Rightarrow> asms" (\<open>_\<close>)
setup \<open>
Document_Output.antiquotation_pretty_source_embedded \<^binding>\<open>const_typ\<close>
--- 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 \<open>Propositional connectives:\<close>
-abbreviation (input) IMPL (infix "impl" 60)
+abbreviation (input) IMPL (infix \<open>impl\<close> 60)
where "\<phi> impl \<psi> \<equiv> \<lambda> xs. \<phi> xs \<longrightarrow> \<psi> xs"
-abbreviation (input) OR (infix "or" 60)
+abbreviation (input) OR (infix \<open>or\<close> 60)
where "\<phi> or \<psi> \<equiv> \<lambda> xs. \<phi> xs \<or> \<psi> xs"
-abbreviation (input) AND (infix "aand" 60)
+abbreviation (input) AND (infix \<open>aand\<close> 60)
where "\<phi> aand \<psi> \<equiv> \<lambda> xs. \<phi> xs \<and> \<psi> xs"
abbreviation (input) not where "not \<phi> \<equiv> \<lambda> xs. \<not> \<phi> xs"
@@ -63,7 +63,7 @@
definition "HLD s = holds (\<lambda>x. x \<in> s)"
-abbreviation HLD_nxt (infixr "\<cdot>" 65) where
+abbreviation HLD_nxt (infixr \<open>\<cdot>\<close> 65) where
"s \<cdot> P \<equiv> HLD s aand nxt P"
context
@@ -79,7 +79,7 @@
alw: "\<lbrakk>\<phi> xs; alw \<phi> (stl xs)\<rbrakk> \<Longrightarrow> alw \<phi> xs"
\<comment> \<open>weak until:\<close>
-coinductive UNTIL (infix "until" 60) for \<phi> \<psi> where
+coinductive UNTIL (infix \<open>until\<close> 60) for \<phi> \<psi> where
base: "\<psi> xs \<Longrightarrow> (\<phi> until \<psi>) xs"
|
step: "\<lbrakk>\<phi> xs; (\<phi> until \<psi>) (stl xs)\<rbrakk> \<Longrightarrow> (\<phi> until \<psi>) xs"
@@ -599,7 +599,7 @@
notes [[inductive_internals]]
begin
-inductive suntil (infix "suntil" 60) for \<phi> \<psi> where
+inductive suntil (infix \<open>suntil\<close> 60) for \<phi> \<psi> where
base: "\<psi> \<omega> \<Longrightarrow> (\<phi> suntil \<psi>) \<omega>"
| step: "\<phi> \<omega> \<Longrightarrow> (\<phi> suntil \<psi>) (stl \<omega>) \<Longrightarrow> (\<phi> suntil \<psi>) \<omega>"
--- 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\<open>Multiplication with a scalar:\<close>
-abbreviation scale :: "('a::times) \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix "*\<^sub>s" 70)
+abbreviation scale :: "('a::times) \<Rightarrow> 'a list \<Rightarrow> 'a list" (infix \<open>*\<^sub>s\<close> 70)
where "x *\<^sub>s xs \<equiv> 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 \<Rightarrow> 'a list \<Rightarrow> 'a" ("\<langle>_,_\<rangle>") where
+definition iprod :: "'a::ring list \<Rightarrow> 'a list \<Rightarrow> 'a" (\<open>\<langle>_,_\<rangle>\<close>) where
"\<langle>xs,ys\<rangle> = (\<Sum>(x,y) \<leftarrow> zip xs ys. x*y)"
lemma iprod_Nil[simp]: "\<langle>[],ys\<rangle> = 0"
--- 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 @@
\<close>
consts
- bind :: "'a \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'd" (infixl "\<bind>" 54)
+ bind :: "'a \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'd" (infixl \<open>\<bind>\<close> 54)
notation (ASCII)
- bind (infixl ">>=" 54)
+ bind (infixl \<open>>>=\<close> 54)
abbreviation (do_notation)
@@ -26,25 +26,25 @@
where "bind_do \<equiv> bind"
notation (output)
- bind_do (infixl "\<bind>" 54)
+ bind_do (infixl \<open>\<bind>\<close> 54)
notation (ASCII output)
- bind_do (infixl ">>=" 54)
+ bind_do (infixl \<open>>>=\<close> 54)
nonterminal do_binds and do_bind
syntax
- "_do_block" :: "do_binds \<Rightarrow> 'a" ("do {//(2 _)//}" [12] 62)
- "_do_bind" :: "[pttrn, 'a] \<Rightarrow> do_bind" ("(2_ \<leftarrow>/ _)" 13)
- "_do_let" :: "[pttrn, 'a] \<Rightarrow> do_bind" ("(2let _ =/ _)" [1000, 13] 13)
- "_do_then" :: "'a \<Rightarrow> do_bind" ("_" [14] 13)
- "_do_final" :: "'a \<Rightarrow> do_binds" ("_")
- "_do_cons" :: "[do_bind, do_binds] \<Rightarrow> do_binds" ("_;//_" [13, 12] 12)
- "_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl "\<then>" 54)
+ "_do_block" :: "do_binds \<Rightarrow> 'a" (\<open>do {//(2 _)//}\<close> [12] 62)
+ "_do_bind" :: "[pttrn, 'a] \<Rightarrow> do_bind" (\<open>(2_ \<leftarrow>/ _)\<close> 13)
+ "_do_let" :: "[pttrn, 'a] \<Rightarrow> do_bind" (\<open>(2let _ =/ _)\<close> [1000, 13] 13)
+ "_do_then" :: "'a \<Rightarrow> do_bind" (\<open>_\<close> [14] 13)
+ "_do_final" :: "'a \<Rightarrow> do_binds" (\<open>_\<close>)
+ "_do_cons" :: "[do_bind, do_binds] \<Rightarrow> do_binds" (\<open>_;//_\<close> [13, 12] 12)
+ "_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl \<open>\<then>\<close> 54)
syntax (ASCII)
- "_do_bind" :: "[pttrn, 'a] \<Rightarrow> do_bind" ("(2_ <-/ _)" 13)
- "_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl ">>" 54)
+ "_do_bind" :: "[pttrn, 'a] \<Rightarrow> do_bind" (\<open>(2_ <-/ _)\<close> 13)
+ "_thenM" :: "['a, 'b] \<Rightarrow> 'c" (infixl \<open>>>\<close> 54)
syntax_consts
"_do_block" "_do_cons" "_do_bind" "_do_then" \<rightleftharpoons> bind and
--- 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 \<Rightarrow> multiset_args" ("_")
- "_multiset_args" :: "'a \<Rightarrow> multiset_args \<Rightarrow> multiset_args" ("_,/ _")
- "_multiset" :: "multiset_args \<Rightarrow> 'a multiset" ("{#(_)#}")
+ "" :: "'a \<Rightarrow> multiset_args" (\<open>_\<close>)
+ "_multiset_args" :: "'a \<Rightarrow> multiset_args \<Rightarrow> multiset_args" (\<open>_,/ _\<close>)
+ "_multiset" :: "multiset_args \<Rightarrow> 'a multiset" (\<open>{#(_)#}\<close>)
syntax_consts
"_multiset_args" "_multiset" == add_mset
translations
@@ -167,12 +167,12 @@
end
syntax
- "_MBall" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool" ("(3\<forall>_\<in>#_./ _)" [0, 0, 10] 10)
- "_MBex" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool" ("(3\<exists>_\<in>#_./ _)" [0, 0, 10] 10)
+ "_MBall" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<forall>_\<in>#_./ _)\<close> [0, 0, 10] 10)
+ "_MBex" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<exists>_\<in>#_./ _)\<close> [0, 0, 10] 10)
syntax (ASCII)
- "_MBall" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool" ("(3\<forall>_:#_./ _)" [0, 0, 10] 10)
- "_MBex" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool" ("(3\<exists>_:#_./ _)" [0, 0, 10] 10)
+ "_MBall" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<forall>_:#_./ _)\<close> [0, 0, 10] 10)
+ "_MBex" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<exists>_:#_./ _)\<close> [0, 0, 10] 10)
syntax_consts
"_MBall" \<rightleftharpoons> Multiset.Ball and
@@ -526,27 +526,27 @@
subsubsection \<open>Pointwise ordering induced by count\<close>
-definition subseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subseteq>#" 50)
+definition subseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix \<open>\<subseteq>#\<close> 50)
where "A \<subseteq># B \<longleftrightarrow> (\<forall>a. count A a \<le> count B a)"
-definition subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subset>#" 50)
+definition subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix \<open>\<subset>#\<close> 50)
where "A \<subset># B \<longleftrightarrow> A \<subseteq># B \<and> A \<noteq> B"
-abbreviation (input) supseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<supseteq>#" 50)
+abbreviation (input) supseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix \<open>\<supseteq>#\<close> 50)
where "supseteq_mset A B \<equiv> B \<subseteq># A"
-abbreviation (input) supset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<supset>#" 50)
+abbreviation (input) supset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix \<open>\<supset>#\<close> 50)
where "supset_mset A B \<equiv> B \<subset># A"
notation (input)
- subseteq_mset (infix "\<le>#" 50) and
- supseteq_mset (infix "\<ge>#" 50)
+ subseteq_mset (infix \<open>\<le>#\<close> 50) and
+ supseteq_mset (infix \<open>\<ge>#\<close> 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 \<open><=#\<close> 50) and
+ subset_mset (infix \<open><#\<close> 50) and
+ supseteq_mset (infix \<open>>=#\<close> 50) and
+ supset_mset (infix \<open>>#\<close> 50)
global_interpretation subset_mset: ordering \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close>
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 \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" ("(1{#_ :# _./ _#})")
+ "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" (\<open>(1{#_ :# _./ _#})\<close>)
syntax
- "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" ("(1{#_ \<in># _./ _#})")
+ "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" (\<open>(1{#_ \<in># _./ _#})\<close>)
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 \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" ("({#_/. _ :# _#})")
+ "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" (\<open>({#_/. _ :# _#})\<close>)
syntax
- "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" ("({#_/. _ \<in># _#})")
+ "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" (\<open>({#_/. _ \<in># _#})\<close>)
syntax_consts
"_comprehension_mset" \<rightleftharpoons> image_mset
translations
"{#e. x \<in># M#}" \<rightleftharpoons> "CONST image_mset (\<lambda>x. e) M"
syntax (ASCII)
- "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" ("({#_/ | _ :# _./ _#})")
+ "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" (\<open>({#_/ | _ :# _./ _#})\<close>)
syntax
- "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" ("({#_/ | _ \<in># _./ _#})")
+ "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" (\<open>({#_/ | _ \<in># _./ _#})\<close>)
syntax_consts
"_comprehension_mset'" \<rightleftharpoons> image_mset
translations
@@ -2684,12 +2684,12 @@
end
-notation sum_mset ("\<Sum>\<^sub>#")
+notation sum_mset (\<open>\<Sum>\<^sub>#\<close>)
syntax (ASCII)
- "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3SUM _:#_. _)" [0, 51, 10] 10)
+ "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" (\<open>(3SUM _:#_. _)\<close> [0, 51, 10] 10)
syntax
- "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3\<Sum>_\<in>#_. _)" [0, 51, 10] 10)
+ "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" (\<open>(3\<Sum>_\<in>#_. _)\<close> [0, 51, 10] 10)
syntax_consts
"_sum_mset_image" \<rightleftharpoons> sum_mset
translations
@@ -2865,12 +2865,12 @@
end
-notation prod_mset ("\<Prod>\<^sub>#")
+notation prod_mset (\<open>\<Prod>\<^sub>#\<close>)
syntax (ASCII)
- "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" ("(3PROD _:#_. _)" [0, 51, 10] 10)
+ "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" (\<open>(3PROD _:#_. _)\<close> [0, 51, 10] 10)
syntax
- "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" ("(3\<Prod>_\<in>#_. _)" [0, 51, 10] 10)
+ "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" (\<open>(3\<Prod>_\<in>#_. _)\<close> [0, 51, 10] 10)
syntax_consts
"_prod_mset_image" \<rightleftharpoons> prod_mset
translations
--- 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>\<open>\<nat>\<close>) This is useful e.g. for the Gamma function.
\<close>
-definition nonpos_Ints ("\<int>\<^sub>\<le>\<^sub>0") where "\<int>\<^sub>\<le>\<^sub>0 = {of_int n |n. n \<le> 0}"
+definition nonpos_Ints (\<open>\<int>\<^sub>\<le>\<^sub>0\<close>) where "\<int>\<^sub>\<le>\<^sub>0 = {of_int n |n. n \<le> 0}"
lemma zero_in_nonpos_Ints [simp,intro]: "0 \<in> \<int>\<^sub>\<le>\<^sub>0"
unfolding nonpos_Ints_def by (auto intro!: exI[of _ "0::int"])
@@ -143,7 +143,7 @@
subsection\<open>Non-negative reals\<close>
-definition nonneg_Reals :: "'a::real_algebra_1 set" ("\<real>\<^sub>\<ge>\<^sub>0")
+definition nonneg_Reals :: "'a::real_algebra_1 set" (\<open>\<real>\<^sub>\<ge>\<^sub>0\<close>)
where "\<real>\<^sub>\<ge>\<^sub>0 = {of_real r | r. r \<ge> 0}"
lemma nonneg_Reals_of_real_iff [simp]: "of_real r \<in> \<real>\<^sub>\<ge>\<^sub>0 \<longleftrightarrow> r \<ge> 0"
@@ -210,7 +210,7 @@
subsection\<open>Non-positive reals\<close>
-definition nonpos_Reals :: "'a::real_algebra_1 set" ("\<real>\<^sub>\<le>\<^sub>0")
+definition nonpos_Reals :: "'a::real_algebra_1 set" (\<open>\<real>\<^sub>\<le>\<^sub>0\<close>)
where "\<real>\<^sub>\<le>\<^sub>0 = {of_real r | r. r \<le> 0}"
lemma nonpos_Reals_of_real_iff [simp]: "of_real r \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> r \<le> 0"
--- 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 \<open>Syntax\<close>
syntax
- "_NumeralType" :: "num_token => type" ("_")
- "_NumeralType0" :: type ("0")
- "_NumeralType1" :: type ("1")
+ "_NumeralType" :: "num_token => type" (\<open>_\<close>)
+ "_NumeralType0" :: type (\<open>0\<close>)
+ "_NumeralType1" :: type (\<open>1\<close>)
syntax_types
"_NumeralType0" == num0 and
--- 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 \<Rightarrow> 'a" ("exec {//(2 _)//}" [12] 62)
- "_sdo_bind" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" ("(_ \<leftarrow>/ _)" 13)
- "_sdo_let" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" ("(2let _ =/ _)" [1000, 13] 13)
- "_sdo_then" :: "'a \<Rightarrow> sdo_bind" ("_" [14] 13)
- "_sdo_final" :: "'a \<Rightarrow> sdo_binds" ("_")
- "_sdo_cons" :: "[sdo_bind, sdo_binds] \<Rightarrow> sdo_binds" ("_;//_" [13, 12] 12)
+ "_sdo_block" :: "sdo_binds \<Rightarrow> 'a" (\<open>exec {//(2 _)//}\<close> [12] 62)
+ "_sdo_bind" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" (\<open>(_ \<leftarrow>/ _)\<close> 13)
+ "_sdo_let" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" (\<open>(2let _ =/ _)\<close> [1000, 13] 13)
+ "_sdo_then" :: "'a \<Rightarrow> sdo_bind" (\<open>_\<close> [14] 13)
+ "_sdo_final" :: "'a \<Rightarrow> sdo_binds" (\<open>_\<close>)
+ "_sdo_cons" :: "[sdo_bind, sdo_binds] \<Rightarrow> sdo_binds" (\<open>_;//_\<close> [13, 12] 12)
syntax (ASCII)
- "_sdo_bind" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" ("(_ <-/ _)" 13)
+ "_sdo_bind" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" (\<open>(_ <-/ _)\<close> 13)
syntax_consts
"_sdo_block" "_sdo_cons" == sdo_syntax and
--- 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 \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "\<^latex>\<open>\\isacharat\<close>" 65)
+ "_appendL" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl \<open>\<^latex>\<open>\isacharat\<close>\<close> 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>\<open>}\\putisatab\\isa{\\ \<close>=\<^latex>\<open>}\\putisatab\\isa{\<close> (_)" [50,49] 50) and
- "Pure.eq" ("(_) \<^latex>\<open>}\\putisatab\\isa{\\ \<close>\<equiv>\<^latex>\<open>}\\putisatab\\isa{\<close> (_)")
+ "HOL.eq" (\<open>(_) \<^latex>\<open>}\putisatab\isa{\ \<close>=\<^latex>\<open>}\putisatab\isa{\<close> (_)\<close> [50,49] 50) and
+ "Pure.eq" (\<open>(_) \<^latex>\<open>}\putisatab\isa{\ \<close>\<equiv>\<^latex>\<open>}\putisatab\isa{\<close> (_)\<close>)
(* 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" (\<open>_::_\<close> [4, 0] 3)
+ "_constrain" :: "prop' => type => prop'" (\<open>_::_\<close> [4, 0] 3)
syntax (output)
- "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3)
- "_constrain" :: "prop' => type => prop'" ("_ :: _" [4, 0] 3)
+ "_constrain" :: "logic => type => logic" (\<open>_ :: _\<close> [4, 0] 3)
+ "_constrain" :: "prop' => type => prop'" (\<open>_ :: _\<close> [4, 0] 3)
(* sorts as intersections *)
syntax (output)
- "_topsort" :: "sort" ("\<top>" 1000)
- "_sort" :: "classes => sort" ("'(_')" 1000)
- "_classes" :: "id => classes => classes" ("_ \<inter> _" 1000)
- "_classes" :: "longid => classes => classes" ("_ \<inter> _" 1000)
+ "_topsort" :: "sort" (\<open>\<top>\<close> 1000)
+ "_sort" :: "classes => sort" (\<open>'(_')\<close> 1000)
+ "_classes" :: "id => classes => classes" (\<open>_ \<inter> _\<close> 1000)
+ "_classes" :: "longid => classes => classes" (\<open>_ \<inter> _\<close> 1000)
end
(*>*)
--- 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 \<open>=:\<close> 1)
declaration \<open>K (Syntax_Phases.term_check 98 "pattern_syntax" (K (map check_pattern_syntax)))\<close>
declaration \<open>K (Syntax_Phases.term_uncheck 98 "pattern_syntax" (map o uncheck_pattern_syntax))\<close>
--- 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 \<circ> phantom = id"
by(simp_all add: o_def id_def)
-syntax "_Phantom" :: "type \<Rightarrow> logic" ("(1Phantom/(1'(_')))")
+syntax "_Phantom" :: "type \<Rightarrow> logic" (\<open>(1Phantom/(1'(_')))\<close>)
syntax_consts "_Phantom" == phantom
translations
"Phantom('t)" => "CONST phantom :: _ \<Rightarrow> ('t, _) phantom"
--- 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 \<Rightarrow> bool \<Rightarrow> 'a" (infixl "when" 20)
+definition "when" :: "'a \<Rightarrow> bool \<Rightarrow> 'a" (infixl \<open>when\<close> 20)
where
"(a when P) = (if P then a else 0)"
@@ -224,7 +224,7 @@
The following type is of central importance:
\<close>
-typedef (overloaded) ('a, 'b) poly_mapping ("(_ \<Rightarrow>\<^sub>0 /_)" [1, 0] 0) =
+typedef (overloaded) ('a, 'b) poly_mapping (\<open>(_ \<Rightarrow>\<^sub>0 /_)\<close> [1, 0] 0) =
"{f :: 'a \<Rightarrow> 'b::zero. finite {x. f x \<noteq> 0}}"
morphisms lookup Abs_poly_mapping
proof -
--- 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 \<longleftrightarrow> x \<le> y \<and> y \<le> x"
notation
- equiv ("'(\<approx>')") and
- equiv ("(_/ \<approx> _)" [51, 51] 50)
+ equiv (\<open>'(\<approx>')\<close>) and
+ equiv (\<open>(_/ \<approx> _)\<close> [51, 51] 50)
lemma equivD1: "x \<le> y" if "x \<approx> y"
using that by (simp add: equiv_def)
--- 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 \<open>OOO\<close> 75) and
+ map_fun (infixr \<open>--->\<close> 55) and
+ rel_fun (infixr \<open>===>\<close> 55)
end
--- 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 @@
\<open>\<sim> :: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close>.\<close>
class eqv =
- fixes eqv :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sim>" 50)
+ fixes eqv :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>\<sim>\<close> 50)
class equiv = eqv +
assumes equiv_refl [intro]: "x \<sim> x"
@@ -76,7 +76,7 @@
text \<open>Abstracted equivalence classes are the canonical representation of
elements of a quotient type.\<close>
-definition "class" :: "'a::equiv \<Rightarrow> 'a quot" ("\<lfloor>_\<rfloor>")
+definition "class" :: "'a::equiv \<Rightarrow> 'a quot" (\<open>\<lfloor>_\<rfloor>\<close>)
where "\<lfloor>a\<rfloor> = Abs_quot {x. a \<sim> x}"
theorem quot_exhaust: "\<exists>a. A = \<lfloor>a\<rfloor>"
--- 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 \<longleftrightarrow> (\<forall>x\<in>set (keys t). x < k)"
-abbreviation rbt_less_symbol (infix "|\<guillemotleft>" 50)
+abbreviation rbt_less_symbol (infix \<open>|\<guillemotleft>\<close> 50)
where "t |\<guillemotleft> x \<equiv> rbt_less x t"
-definition rbt_greater :: "'a \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool" (infix "\<guillemotleft>|" 50)
+definition rbt_greater :: "'a \<Rightarrow> ('a, 'b) rbt \<Rightarrow> bool" (infix \<open>\<guillemotleft>|\<close> 50)
where
rbt_greater_prop: "rbt_greater k t = (\<forall>x\<in>set (keys t). k < x)"
--- 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 \<open>The $n$-element subsets of a set $A$\<close>
-definition nsets :: "['a set, nat] \<Rightarrow> 'a set set" ("([_]\<^bsup>_\<^esup>)" [0,999] 999)
+definition nsets :: "['a set, nat] \<Rightarrow> 'a set set" (\<open>([_]\<^bsup>_\<^esup>)\<close> [0,999] 999)
where "nsets A n \<equiv> {N. N \<subseteq> A \<and> finite N \<and> card N = n}"
lemma finite_imp_finite_nsets: "finite A \<Longrightarrow> finite ([A]\<^bsup>k\<^esup>)"
--- 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 \<Rightarrow> real \<Rightarrow> real" (infixl "rmod" 70) where
+definition rmod :: "real \<Rightarrow> real \<Rightarrow> real" (infixl \<open>rmod\<close> 70) where
"x rmod y = x - \<bar>y\<bar> * of_int \<lfloor>x / \<bar>y\<bar>\<rfloor>"
lemma rmod_conv_frac: "y \<noteq> 0 \<Longrightarrow> x rmod y = frac (x / \<bar>y\<bar>) * \<bar>y\<bar>"
--- 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::{}" ("\<hole>")
+consts rewrite_HOLE :: "'a::{}" (\<open>\<hole>\<close>)
lemma eta_expand:
fixes f :: "'a::{} \<Rightarrow> 'b::{}"
--- 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 \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "+o" 70)
+definition elt_set_plus :: "'a::plus \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl \<open>+o\<close> 70)
where "a +o B = {c. \<exists>b\<in>B. c = a + b}"
-definition elt_set_times :: "'a::times \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "*o" 80)
+definition elt_set_times :: "'a::times \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl \<open>*o\<close> 80)
where "a *o B = {c. \<exists>b\<in>B. c = a * b}"
-abbreviation (input) elt_set_eq :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infix "=o" 50)
+abbreviation (input) elt_set_eq :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infix \<open>=o\<close> 50)
where "x =o A \<equiv> x \<in> A"
instance set :: (semigroup_add) semigroup_add
--- 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\<open>Idioms for being a suitable union/intersection of something\<close>
definition union_of :: "('a set set \<Rightarrow> bool) \<Rightarrow> ('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
- (infixr "union'_of" 60)
+ (infixr \<open>union'_of\<close> 60)
where "P union_of Q \<equiv> \<lambda>S. \<exists>\<U>. P \<U> \<and> \<U> \<subseteq> Collect Q \<and> \<Union>\<U> = S"
definition intersection_of :: "('a set set \<Rightarrow> bool) \<Rightarrow> ('a set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
- (infixr "intersection'_of" 60)
+ (infixr \<open>intersection'_of\<close> 60)
where "P intersection_of Q \<equiv> \<lambda>S. \<exists>\<U>. P \<U> \<and> \<U> \<subseteq> Collect Q \<and> \<Inter>\<U> = S"
definition arbitrary:: "'a set set \<Rightarrow> bool" where "arbitrary \<U> \<equiv> True"
@@ -310,7 +310,7 @@
text\<open>A somewhat cheap but handy way of getting localized forms of various topological concepts
(open, closed, borel, fsigma, gdelta etc.)\<close>
-definition relative_to :: "['a set \<Rightarrow> bool, 'a set, 'a set] \<Rightarrow> bool" (infixl "relative'_to" 55)
+definition relative_to :: "['a set \<Rightarrow> bool, 'a set, 'a set] \<Rightarrow> bool" (infixl \<open>relative'_to\<close> 55)
where "P relative_to S \<equiv> \<lambda>T. \<exists>U. P U \<and> S \<inter> U = T"
lemma relative_to_UNIV [simp]: "(P relative_to UNIV) S \<longleftrightarrow> P S"
--- 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 \<open>Parallel lists\<close>
-definition parallel :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infixl "\<parallel>" 50)
+definition parallel :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infixl \<open>\<parallel>\<close> 50)
where "(xs \<parallel> ys) = (\<not> prefix xs ys \<and> \<not> prefix ys xs)"
lemma parallelI [intro]: "\<not> prefix xs ys \<Longrightarrow> \<not> prefix ys xs \<Longrightarrow> xs \<parallel> ys"
--- 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 ("\<langle>\<rangle>") |
- Node "'a tree" ("value": 'a) "'a tree" ("(1\<langle>_,/ _,/ _\<rangle>)")
+ Leaf (\<open>\<langle>\<rangle>\<close>) |
+ Node "'a tree" ("value": 'a) "'a tree" (\<open>(1\<langle>_,/ _,/ _\<rangle>)\<close>)
datatype_compat tree
primrec left :: "'a tree \<Rightarrow> 'a tree" where
--- 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 ("'(\<le>s')") and
- word_sle ("(_/ \<le>s _)" [51, 51] 50) and
- word_sless ("'(<s')") and
- word_sless ("(_/ <s _)" [51, 51] 50)
+ word_sle (\<open>'(\<le>s')\<close>) and
+ word_sle (\<open>(_/ \<le>s _)\<close> [51, 51] 50) and
+ word_sless (\<open>'(<s')\<close>) and
+ word_sless (\<open>(_/ <s _)\<close> [51, 51] 50)
notation (input)
- word_sle ("(_/ <=s _)" [51, 51] 50)
+ word_sle (\<open>(_/ <=s _)\<close> [51, 51] 50)
lemma word_sle_eq [code]:
\<open>a <=s b \<longleftrightarrow> sint a \<le> sint b\<close>
--- 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 \<open>Definitions\<close>
-definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" ("(1O'(_'))") where
+definition bigo :: "('a => 'b::linordered_idom) => ('a => 'b) set" (\<open>(1O'(_'))\<close>) where
"O(f::('a => 'b)) == {h. \<exists>c. \<forall>x. \<bar>h x\<bar> <= c * \<bar>f x\<bar>}"
lemma bigo_pos_const:
@@ -610,7 +610,7 @@
subsection \<open>Less than or equal to\<close>
-definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "<o" 70) where
+definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl \<open><o\<close> 70) where
"f <o g == (\<lambda>x. max (f x - g x) 0)"
lemma bigo_lesseq1: "f =o O(h) \<Longrightarrow> \<forall>x. \<bar>g x\<bar> <= \<bar>f x\<bar> \<Longrightarrow>
--- 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\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
nonterminal mtuple_args
syntax
- "" :: "'a \<Rightarrow> mtuple_args" ("_")
- "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" ("_,/ _")
- "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)")
+ "" :: "'a \<Rightarrow> mtuple_args" (\<open>_\<close>)
+ "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" (\<open>_,/ _\<close>)
+ "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" (\<open>(2\<lbrace>_,/ _\<rbrace>)\<close>)
syntax_consts
"_MTuple_args" "_MTuple" \<rightleftharpoons> MPair
translations
@@ -60,7 +60,7 @@
"\<lbrace>x, y\<rbrace>" \<rightleftharpoons> "CONST MPair x y"
-definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where
+definition HPair :: "[msg,msg] => msg" (\<open>(4Hash[_] /_)\<close> [0, 1000]) where
\<comment> \<open>Message Y paired with a MAC computed with the help of X\<close>
"Hash[X] Y == \<lbrace> Hash\<lbrace>X,Y\<rbrace>, Y\<rbrace>"
--- 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 @@
\<lparr>pset = S, order = induced S (order cl)\<rparr> \<in> CompleteLattice}"
abbreviation
- sublattice_syntax :: "['a set, 'a potype] => bool" ("_ <<= _" [51, 50] 50)
+ sublattice_syntax :: "['a set, 'a potype] => bool" (\<open>_ <<= _\<close> [51, 50] 50)
where "S <<= cl \<equiv> S \<in> sublattice `` {cl}"
definition dual :: "'a potype => 'a potype" where
--- 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"}
\<close>
-definition intervall :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" ("_ \<in> [_, _')") where
+definition intervall :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" (\<open>_ \<in> [_, _')\<close>) where
"x \<in> [a, b) \<equiv> a \<le> x \<and> x < b"
lemma pc_0: "x < n \<Longrightarrow> (x \<in> [0, n) \<Longrightarrow> P x) \<Longrightarrow> 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 ("\<phi>\<^sub>a") where
+definition phi_append :: method_type (\<open>\<phi>\<^sub>a\<close>) where
"\<phi>\<^sub>a \<equiv> map (\<lambda>(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 ("\<phi>\<^sub>m") where
+definition phi_makelist :: method_type (\<open>\<phi>\<^sub>m\<close>) where
"\<phi>\<^sub>m \<equiv> map (\<lambda>(x,y). Some (x, y)) [
( [], [OK Ctest, Err , Err ]),
( [Clist], [OK Ctest, Err , Err ]),
@@ -369,7 +369,7 @@
done
text \<open>The whole program is welltyped:\<close>
-definition Phi :: prog_type ("\<Phi>") where
+definition Phi :: prog_type (\<open>\<Phi>\<close>) where
"\<Phi> C sg \<equiv> if C = test_name \<and> sg = (makelist_name, []) then \<phi>\<^sub>m else
if C = list_name \<and> sg = (append_name, [Class list_name]) then \<phi>\<^sub>a else []"
--- 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 ("\<Gamma>") and Phi ("\<Phi>")
+ fixes G (\<open>\<Gamma>\<close>) and Phi (\<open>\<Phi>\<close>)
assumes wt: "wt_jvm_prog \<Gamma> \<Phi>"
assumes is_class: "is_class \<Gamma> C"
and "method": "method (\<Gamma>,C) (m,[]) = Some (C, b)"
@@ -447,7 +447,7 @@
state or in the canonical start state)
\<close>
corollary welltyped_commutes:
- fixes G ("\<Gamma>") and Phi ("\<Phi>")
+ fixes G (\<open>\<Gamma>\<close>) and Phi (\<open>\<Phi>\<close>)
assumes wt: "wt_jvm_prog \<Gamma> \<Phi>" and *: "\<Gamma>,\<Phi> \<turnstile>JVM s \<surd>"
shows "\<Gamma> \<turnstile> (Normal s) \<midarrow>jvmd\<rightarrow> (Normal t) = \<Gamma> \<turnstile> s \<midarrow>jvm\<rightarrow> t"
apply rule
@@ -458,7 +458,7 @@
done
corollary welltyped_initial_commutes:
- fixes G ("\<Gamma>") and Phi ("\<Phi>")
+ fixes G (\<open>\<Gamma>\<close>) and Phi (\<open>\<Phi>\<close>)
assumes wt: "wt_jvm_prog \<Gamma> \<Phi>"
assumes is_class: "is_class \<Gamma> C"
and "method": "method (\<Gamma>,C) (m,[]) = Some (C, b)"
--- 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 ("\<Gamma>")
+ fixes G :: jvm_prog (\<open>\<Gamma>\<close>)
assumes wf: "wf_prog wf_mb \<Gamma>"
shows hconf_start: "\<Gamma> \<turnstile>h (start_heap \<Gamma>) \<surd>"
apply (unfold hconf_def start_heap_def)
@@ -1330,7 +1330,7 @@
done
lemma
- fixes G :: jvm_prog ("\<Gamma>") and Phi :: prog_type ("\<Phi>")
+ fixes G :: jvm_prog (\<open>\<Gamma>\<close>) and Phi :: prog_type (\<open>\<Phi>\<close>)
shows BV_correct_initial:
"wt_jvm_prog \<Gamma> \<Phi> \<Longrightarrow> is_class \<Gamma> C \<Longrightarrow> method (\<Gamma>,C) (m,[]) = Some (C, b)
\<Longrightarrow> \<Gamma>,\<Phi> \<turnstile>JVM start_state G C m \<surd>"
@@ -1347,8 +1347,8 @@
done
theorem typesafe:
- fixes G :: jvm_prog ("\<Gamma>")
- and Phi :: prog_type ("\<Phi>")
+ fixes G :: jvm_prog (\<open>\<Gamma>\<close>)
+ and Phi :: prog_type (\<open>\<Phi>\<close>)
assumes welltyped: "wt_jvm_prog \<Gamma> \<Phi>"
and main_method: "is_class \<Gamma> C" "method (\<Gamma>,C) (m,[]) = Some (C, b)"
and exec_all: "G \<turnstile> start_state \<Gamma> C m \<midarrow>jvm\<rightarrow> s"
--- 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] \<Rightarrow> bool"
- ("_,_ \<turnstile>JVM _ \<surd>" [51,51] 50) where
+ (\<open>_,_ \<turnstile>JVM _ \<surd>\<close> [51,51] 50) where
"correct_state G phi == \<lambda>(xp,hp,frs).
case xp of
None \<Rightarrow> (case frs of
--- 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] \<Rightarrow> bool"
- ("_ \<turnstile> _ <=o _" [71,71] 70) where
+ (\<open>_ \<turnstile> _ <=o _\<close> [71,71] 70) where
"sup_ty_opt G == Err.le (subtype G)"
definition sup_loc :: "['code prog,locvars_type,locvars_type] \<Rightarrow> bool"
- ("_ \<turnstile> _ <=l _" [71,71] 70) where
+ (\<open>_ \<turnstile> _ <=l _\<close> [71,71] 70) where
"sup_loc G == Listn.le (sup_ty_opt G)"
definition sup_state :: "['code prog,state_type,state_type] \<Rightarrow> bool"
- ("_ \<turnstile> _ <=s _" [71,71] 70) where
+ (\<open>_ \<turnstile> _ <=s _\<close> [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] \<Rightarrow> bool"
- ("_ \<turnstile> _ <='' _" [71,71] 70) where
+ (\<open>_ \<turnstile> _ <='' _\<close> [71,71] 70) where
"sup_state_opt G == Opt.le (sup_state G)"
--- 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 \<Rightarrow>
aheap \<Rightarrow> opstack \<Rightarrow> locvars \<Rightarrow>
bool"
- ("{_,_,_} \<turnstile> {_, _, _} >- _ \<rightarrow> {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60) where
+ (\<open>{_,_,_} \<turnstile> {_, _, _} >- _ \<rightarrow> {_, _, _}\<close> [61,61,61,61,61,61,90,61,61,61]60) where
"{G,C,S} \<turnstile> {hp0, os0, lvars0} >- instrs \<rightarrow> {hp1, os1, lvars1} ==
\<forall>pre post frs.
(gis (gmb G C S) = pre @ instrs @ post) \<longrightarrow>
--- 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 \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd"
- (infixr "\<box>" 55)
+ (infixr \<open>\<box>\<close> 55)
where
"comb == (\<lambda> f1 f2 x0. let (xs1, x1) = f1 x0;
(xs2, x2) = f2 x1
--- 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" ("\<phi>")
+ fixes phi :: "'a list" (\<open>\<phi>\<close>)
fixes c :: "'a list"
defines cert_def: "c \<equiv> make_cert step \<phi> \<bottom>"
--- 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 (\<open>s\<^sub>0\<close>)
fixes c :: "'a list"
fixes ins :: "'b list"
- fixes phi :: "'a list" ("\<phi>")
+ fixes phi :: "'a list" (\<open>\<phi>\<close>)
defines phi_def:
"\<phi> \<equiv> map (\<lambda>pc. if c!pc = \<bottom> then wtl (take pc ins) c 0 s0 else c!pc)
[0..<length ins]"
--- a/src/HOL/MicroJava/DFA/LBVSpec.thy Fri Sep 20 19:07:10 2024 +0200
+++ b/src/HOL/MicroJava/DFA/LBVSpec.thy Fri Sep 20 19:51:08 2024 +0200
@@ -44,8 +44,8 @@
"bottom r B \<equiv> \<forall>x. B <=_r x"
locale lbv = Semilat +
- fixes T :: "'a" ("\<top>")
- fixes B :: "'a" ("\<bottom>")
+ fixes T :: "'a" (\<open>\<top>\<close>)
+ fixes B :: "'a" (\<open>\<bottom>\<close>)
fixes step :: "'a step_type"
assumes top: "top r \<top>"
assumes T_A: "\<top> \<in> A"
--- a/src/HOL/MicroJava/DFA/Listn.thy Fri Sep 20 19:07:10 2024 +0200
+++ b/src/HOL/MicroJava/DFA/Listn.thy Fri Sep 20 19:51:08 2024 +0200
@@ -17,12 +17,12 @@
abbreviation
lesublist_syntax :: "'a list \<Rightarrow> 'a ord \<Rightarrow> 'a list \<Rightarrow> bool"
- ("(_ /<=[_] _)" [50, 0, 51] 50)
+ (\<open>(_ /<=[_] _)\<close> [50, 0, 51] 50)
where "x <=[r] y == x <=_(le r) y"
abbreviation
lesssublist_syntax :: "'a list \<Rightarrow> 'a ord \<Rightarrow> 'a list \<Rightarrow> bool"
- ("(_ /<[_] _)" [50, 0, 51] 50)
+ (\<open>(_ /<[_] _)\<close> [50, 0, 51] 50)
where "x <[r] y == x <_(le r) y"
definition map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list" where
@@ -30,7 +30,7 @@
abbreviation
plussublist_syntax :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b list \<Rightarrow> 'c list"
- ("(_ /+[_] _)" [65, 0, 66] 65)
+ (\<open>(_ /+[_] _)\<close> [65, 0, 66] 65)
where "x +[f] y == x +_(map2 f) y"
primrec coalesce :: "'a err list \<Rightarrow> 'a list err" where
--- 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 \<Rightarrow> 'a ord \<Rightarrow> 'b ord \<Rightarrow> 'a * 'b \<Rightarrow> bool"
- ("(_ /<='(_,_') _)" [50, 0, 0, 51] 50)
+ (\<open>(_ /<='(_,_') _)\<close> [50, 0, 0, 51] 50)
where "p <=(rA,rB) q == p <=_(le rA rB) q"
lemma unfold_lesub_prod:
--- 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" (\<open>(_ /<='__ _)\<close> [50, 1000, 51] 50) and
+ "lesssub" (\<open>(_ /<'__ _)\<close> [50, 1000, 51] 50) and
+ "plussub" (\<open>(_ /+'__ _)\<close> [65, 1000, 66] 65)
notation
- "lesub" ("(_ /\<sqsubseteq>\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and
- "lesssub" ("(_ /\<sqsubset>\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and
- "plussub" ("(_ /\<squnion>\<^bsub>_\<^esub> _)" [65, 0, 66] 65)
+ "lesub" (\<open>(_ /\<sqsubseteq>\<^bsub>_\<^esub> _)\<close> [50, 0, 51] 50) and
+ "lesssub" (\<open>(_ /\<sqsubset>\<^bsub>_\<^esub> _)\<close> [50, 0, 51] 50) and
+ "plussub" (\<open>(_ /\<squnion>\<^bsub>_\<^esub> _)\<close> [65, 0, 66] 65)
(* allow \<sub> instead of \<bsub>..\<esub> *)
abbreviation (input)
- lesub1 :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubseteq>\<^sub>_ _)" [50, 1000, 51] 50)
+ lesub1 :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" (\<open>(_ /\<sqsubseteq>\<^sub>_ _)\<close> [50, 1000, 51] 50)
where "x \<sqsubseteq>\<^sub>r y == x \<sqsubseteq>\<^bsub>r\<^esub> y"
abbreviation (input)
- lesssub1 :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubset>\<^sub>_ _)" [50, 1000, 51] 50)
+ lesssub1 :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" (\<open>(_ /\<sqsubset>\<^sub>_ _)\<close> [50, 1000, 51] 50)
where "x \<sqsubset>\<^sub>r y == x \<sqsubset>\<^bsub>r\<^esub> y"
abbreviation (input)
- plussub1 :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /\<squnion>\<^sub>_ _)" [65, 1000, 66] 65)
+ plussub1 :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" (\<open>(_ /\<squnion>\<^sub>_ _)\<close> [65, 1000, 66] 65)
where "x \<squnion>\<^sub>f y == x \<squnion>\<^bsub>f\<^esub> y"
definition ord :: "('a \<times> 'a) set \<Rightarrow> 'a ord" where
--- 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 \<times> 's) list \<Rightarrow> 's ord \<Rightarrow> (nat \<times> 's) list \<Rightarrow> bool"
- ("(_ /\<le>|_| _)" [50, 0, 51] 50) where
+ (\<open>(_ /\<le>|_| _)\<close> [50, 0, 51] 50) where
"x \<le>|r| y \<equiv> \<forall>(p,s) \<in> set x. \<exists>s'. (p,s') \<in> set y \<and> s <=_r s'"
-primrec plusplussub :: "'a list \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" ("(_ /++'__ _)" [65, 1000, 66] 65) where
+primrec plusplussub :: "'a list \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>(_ /++'__ _)\<close> [65, 1000, 66] 65) where
"[] ++_f y = y"
| "(x#xs) ++_f y = xs ++_f (x +_f y)"
--- 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 \<times> (vname \<rightharpoonup> ty)" \<comment> \<open>same as \<open>env\<close> of \<open>WellType.thy\<close>\<close>
-definition hext :: "aheap => aheap => bool" ("_ \<le>| _" [51,51] 50) where
+definition hext :: "aheap => aheap => bool" (\<open>_ \<le>| _\<close> [51,51] 50) where
"h\<le>|h' == \<forall>a C fs. h a = Some(C,fs) --> (\<exists>fs'. h' a = Some(C,fs'))"
definition conf :: "'c prog => aheap => val => ty => bool"
- ("_,_ \<turnstile> _ ::\<preceq> _" [51,51,51,51] 50) where
+ (\<open>_,_ \<turnstile> _ ::\<preceq> _\<close> [51,51,51,51] 50) where
"G,h\<turnstile>v::\<preceq>T == \<exists>T'. typeof (map_option obj_ty o h) v = Some T' \<and> G\<turnstile>T'\<preceq>T"
definition lconf :: "'c prog => aheap => ('a \<rightharpoonup> val) => ('a \<rightharpoonup> ty) => bool"
- ("_,_ \<turnstile> _ [::\<preceq>] _" [51,51,51,51] 50) where
+ (\<open>_,_ \<turnstile> _ [::\<preceq>] _\<close> [51,51,51,51] 50) where
"G,h\<turnstile>vs[::\<preceq>]Ts == \<forall>n T. Ts n = Some T --> (\<exists>v. vs n = Some v \<and> G,h\<turnstile>v::\<preceq>T)"
-definition oconf :: "'c prog => aheap => obj => bool" ("_,_ \<turnstile> _ \<surd>" [51,51,51] 50) where
+definition oconf :: "'c prog => aheap => obj => bool" (\<open>_,_ \<turnstile> _ \<surd>\<close> [51,51,51] 50) where
"G,h\<turnstile>obj \<surd> == G,h\<turnstile>snd obj[::\<preceq>]map_of (fields (G,fst obj))"
-definition hconf :: "'c prog => aheap => bool" ("_ \<turnstile>h _ \<surd>" [51,51] 50) where
+definition hconf :: "'c prog => aheap => bool" (\<open>_ \<turnstile>h _ \<surd>\<close> [51,51] 50) where
"G\<turnstile>h h \<surd> == \<forall>a obj. h a = Some obj --> G,h\<turnstile>obj \<surd>"
definition xconf :: "aheap \<Rightarrow> val option \<Rightarrow> bool" where
"xconf hp vo == preallocated hp \<and> (\<forall> v. (vo = Some v) \<longrightarrow> (\<exists> xc. v = (Addr (XcptRef xc))))"
-definition conforms :: "xstate => java_mb env' => bool" ("_ ::\<preceq> _" [51,51] 50) where
+definition conforms :: "xstate => java_mb env' => bool" (\<open>_ ::\<preceq> _\<close> [51,51] 50) where
"s::\<preceq>E == prg E\<turnstile>h heap (store s) \<surd> \<and>
prg E,heap (store s)\<turnstile>locals (store s)[::\<preceq>]localT E \<and>
xconf (heap (store s)) (abrupt s)"
--- 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 @@
\<comment> \<open>Auxiliary notions\<close>
-definition fits :: "java_mb prog \<Rightarrow> state \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) where
+definition fits :: "java_mb prog \<Rightarrow> state \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" (\<open>_,_\<turnstile>_ fits _\<close>[61,61,61,61]60) where
"G,s\<turnstile>a' fits T \<equiv> case T of PrimT T' \<Rightarrow> False | RefT T' \<Rightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
-definition catch :: "java_mb prog \<Rightarrow> xstate \<Rightarrow> cname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60) where
+definition catch :: "java_mb prog \<Rightarrow> xstate \<Rightarrow> cname \<Rightarrow> bool" (\<open>_,_\<turnstile>catch _\<close>[61,61,61]60) where
"G,s\<turnstile>catch C\<equiv> case abrupt s of None \<Rightarrow> False | Some a \<Rightarrow> G,store s\<turnstile> a fits Class C"
-definition lupd :: "vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state" ("lupd'(_\<mapsto>_')"[10,10]1000) where
+definition lupd :: "vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state" (\<open>lupd'(_\<mapsto>_')\<close>[10,10]1000) where
"lupd vn v \<equiv> \<lambda> (hp,loc). (hp, (loc(vn\<mapsto>v)))"
definition new_xcpt_var :: "vname \<Rightarrow> xstate \<Rightarrow> xstate" where
@@ -27,12 +27,12 @@
inductive
eval :: "[java_mb prog,xstate,expr,val,xstate] => bool "
- ("_ \<turnstile> _ -_\<succ>_-> _" [51,82,60,82,82] 81)
+ (\<open>_ \<turnstile> _ -_\<succ>_-> _\<close> [51,82,60,82,82] 81)
and evals :: "[java_mb prog,xstate,expr list,
val list,xstate] => bool "
- ("_ \<turnstile> _ -_[\<succ>]_-> _" [51,82,60,51,82] 81)
+ (\<open>_ \<turnstile> _ -_[\<succ>]_-> _\<close> [51,82,60,51,82] 81)
and exec :: "[java_mb prog,xstate,stmt, xstate] => bool "
- ("_ \<turnstile> _ -_-> _" [51,82,60,82] 81)
+ (\<open>_ \<turnstile> _ -_-> _\<close> [51,82,60,82] 81)
for G :: "java_mb prog"
where
--- 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 \<comment> \<open>literal value, also references\<close>
| BinOp binop expr expr \<comment> \<open>binary operation\<close>
| LAcc vname \<comment> \<open>local (incl. parameter) access\<close>
- | LAss vname expr ("_::=_" [90,90]90) \<comment> \<open>local assign\<close>
- | FAcc cname expr vname ("{_}_.._" [10,90,99]90) \<comment> \<open>field access\<close>
+ | LAss vname expr (\<open>_::=_\<close> [90,90]90) \<comment> \<open>local assign\<close>
+ | FAcc cname expr vname (\<open>{_}_.._\<close> [10,90,99]90) \<comment> \<open>field access\<close>
| FAss cname expr vname
- expr ("{_}_.._:=_" [10,90,99,90]90) \<comment> \<open>field ass.\<close>
+ expr (\<open>{_}_.._:=_\<close> [10,90,99,90]90) \<comment> \<open>field ass.\<close>
| Call cname expr mname
- "ty list" "expr list" ("{_}_.._'( {_}_')" [10,90,99,10,10] 90) \<comment> \<open>method call\<close>
+ "ty list" "expr list" (\<open>{_}_.._'( {_}_')\<close> [10,90,99,10,10] 90) \<comment> \<open>method call\<close>
datatype_compat expr
datatype stmt
= Skip \<comment> \<open>empty statement\<close>
| Expr expr \<comment> \<open>expression statement\<close>
- | 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 (\<open>_;; _\<close> [61,60]60)
+ | Cond expr stmt stmt (\<open>If '(_') _ Else _\<close> [80,79,79]70)
+ | Loop expr stmt (\<open>While '(_') _\<close> [80,79]70)
end
--- 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 \<times> cname) set"
- and subcls1' :: "'c prog => cname \<Rightarrow> cname => bool" ("_ \<turnstile> _ \<prec>C1 _" [71,71,71] 70)
+ and subcls1' :: "'c prog => cname \<Rightarrow> cname => bool" (\<open>_ \<turnstile> _ \<prec>C1 _\<close> [71,71,71] 70)
for G :: "'c prog"
where
"G \<turnstile> C \<prec>C1 D \<equiv> (C, D) \<in> subcls1 G"
| subcls1I: "\<lbrakk>class G C = Some (D,rest); C \<noteq> Object\<rbrakk> \<Longrightarrow> G \<turnstile> C \<prec>C1 D"
abbreviation
- subcls :: "'c prog => cname \<Rightarrow> cname => bool" ("_ \<turnstile> _ \<preceq>C _" [71,71,71] 70)
+ subcls :: "'c prog => cname \<Rightarrow> cname => bool" (\<open>_ \<turnstile> _ \<preceq>C _\<close> [71,71,71] 70)
where "G \<turnstile> C \<preceq>C D \<equiv> (C, D) \<in> (subcls1 G)\<^sup>*"
lemma subcls1D:
@@ -217,7 +217,7 @@
\<comment> \<open>widening, viz. method invocation conversion,cf. 5.3 i.e. sort of syntactic subtyping\<close>
inductive
- widen :: "'c prog => [ty , ty ] => bool" ("_ \<turnstile> _ \<preceq> _" [71,71,71] 70)
+ widen :: "'c prog => [ty , ty ] => bool" (\<open>_ \<turnstile> _ \<preceq> _\<close> [71,71,71] 70)
for G :: "'c prog"
where
refl [intro!, simp]: "G\<turnstile> T \<preceq> T" \<comment> \<open>identity conv., cf. 5.1.1\<close>
@@ -231,7 +231,7 @@
\<comment> \<open>casting conversion, cf. 5.5 / 5.1.5\<close>
\<comment> \<open>left out casts on primitve types\<close>
inductive
- cast :: "'c prog => [ty , ty ] => bool" ("_ \<turnstile> _ \<preceq>? _" [71,71,71] 70)
+ cast :: "'c prog => [ty , ty ] => bool" (\<open>_ \<turnstile> _ \<preceq>? _\<close> [71,71,71] 70)
for G :: "'c prog"
where
widen: "G\<turnstile> C\<preceq> D ==> G\<turnstile>C \<preceq>? D"
--- 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 @@
\<comment> \<open>local variables might include This, which is hidden anyway\<close>
inductive
- ty_expr :: "'c env => expr => ty => bool" ("_ \<turnstile> _ :: _" [51, 51, 51] 50)
- and ty_exprs :: "'c env => expr list => ty list => bool" ("_ \<turnstile> _ [::] _" [51, 51, 51] 50)
- and wt_stmt :: "'c env => stmt => bool" ("_ \<turnstile> _ \<surd>" [51, 51] 50)
+ ty_expr :: "'c env => expr => ty => bool" (\<open>_ \<turnstile> _ :: _\<close> [51, 51, 51] 50)
+ and ty_exprs :: "'c env => expr list => ty list => bool" (\<open>_ \<turnstile> _ [::] _\<close> [51, 51, 51] 50)
+ and wt_stmt :: "'c env => stmt => bool" (\<open>_ \<turnstile> _ \<surd>\<close> [51, 51] 50)
where
NewC: "[| is_class (prg E) C |] ==>
--- 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 \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool"
- ("_ \<turnstile> _ \<midarrow>jvmd\<rightarrow> _" [61,61,61]60) where
+ (\<open>_ \<turnstile> _ \<midarrow>jvmd\<rightarrow> _\<close> [61,61,61]60) where
"G \<turnstile> s \<midarrow>jvmd\<rightarrow> t \<longleftrightarrow>
(s,t) \<in> ({(s,t). exec_d G s = TypeError \<and> t = TypeError} \<union>
{(s,t). \<exists>t'. exec_d G s = Normal (Some t') \<and> t = Normal t'})\<^sup>*"
--- 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"
- ("_ \<turnstile> _ \<midarrow>jvm\<rightarrow> _" [61,61,61]60) where
+ (\<open>_ \<turnstile> _ \<midarrow>jvm\<rightarrow> _\<close> [61,61,61]60) where
"G \<turnstile> s \<midarrow>jvm\<rightarrow> t == (s,t) \<in> {(s,t). exec(G,s) = Some t}\<^sup>*"
--- 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" ("_ |\<turnstile>/ _" [61, 61] 60)
- and ehoare :: "[triple set, etriple] => bool" ("_ |\<turnstile>\<^sub>e/ _" [61, 61] 60)
+ hoare :: "[triple set, triple set] => bool" (\<open>_ |\<turnstile>/ _\<close> [61, 61] 60)
+ and ehoare :: "[triple set, etriple] => bool" (\<open>_ |\<turnstile>\<^sub>e/ _\<close> [61, 61] 60)
and hoare1 :: "[triple set, assn,stmt,assn] => bool"
- ("_ \<turnstile>/ ({(1_)}/ (_)/ {(1_)})" [61, 3, 90, 3] 60)
+ (\<open>_ \<turnstile>/ ({(1_)}/ (_)/ {(1_)})\<close> [61, 3, 90, 3] 60)
and ehoare1 :: "[triple set, assn,expr,vassn]=> bool"
- ("_ \<turnstile>\<^sub>e/ ({(1_)}/ (_)/ {(1_)})" [61, 3, 90, 3] 60)
+ (\<open>_ \<turnstile>\<^sub>e/ ({(1_)}/ (_)/ {(1_)})\<close> [61, 3, 90, 3] 60)
where
"A \<turnstile> {P}c{Q} \<equiv> A |\<turnstile> {(P,c,Q)}"
--- 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" ("\<Turnstile> {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) where
+definition valid :: "[assn,stmt, assn] => bool" (\<open>\<Turnstile> {(1_)}/ (_)/ {(1_)}\<close> [3,90,3] 60) where
"\<Turnstile> {P} c {Q} \<equiv> \<forall>s t. P s --> (\<exists>n. s -c -n\<rightarrow> t) --> Q t"
-definition evalid :: "[assn,expr,vassn] => bool" ("\<Turnstile>\<^sub>e {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) where
+definition evalid :: "[assn,expr,vassn] => bool" (\<open>\<Turnstile>\<^sub>e {(1_)}/ (_)/ {(1_)}\<close> [3,90,3] 60) where
"\<Turnstile>\<^sub>e {P} e {Q} \<equiv> \<forall>s v t. P s --> (\<exists>n. s -e\<succ>v-n\<rightarrow> t) --> Q v t"
-definition nvalid :: "[nat, triple ] => bool" ("\<Turnstile>_: _" [61,61] 60) where
+definition nvalid :: "[nat, triple ] => bool" (\<open>\<Turnstile>_: _\<close> [61,61] 60) where
"\<Turnstile>n: t \<equiv> let (P,c,Q) = t in \<forall>s t. s -c -n\<rightarrow> t --> P s --> Q t"
-definition envalid :: "[nat,etriple ] => bool" ("\<Turnstile>_:\<^sub>e _" [61,61] 60) where
+definition envalid :: "[nat,etriple ] => bool" (\<open>\<Turnstile>_:\<^sub>e _\<close> [61,61] 60) where
"\<Turnstile>n:\<^sub>e t \<equiv> let (P,e,Q) = t in \<forall>s v t. s -e\<succ>v-n\<rightarrow> t --> P s --> Q v t"
-definition nvalids :: "[nat, triple set] => bool" ("|\<Turnstile>_: _" [61,61] 60) where
+definition nvalids :: "[nat, triple set] => bool" (\<open>|\<Turnstile>_: _\<close> [61,61] 60) where
"|\<Turnstile>n: T \<equiv> \<forall>t\<in>T. \<Turnstile>n: t"
-definition cnvalids :: "[triple set,triple set] => bool" ("_ |\<Turnstile>/ _" [61,61] 60) where
+definition cnvalids :: "[triple set,triple set] => bool" (\<open>_ |\<Turnstile>/ _\<close> [61,61] 60) where
"A |\<Turnstile> C \<equiv> \<forall>n. |\<Turnstile>n: A --> |\<Turnstile>n: C"
-definition cenvalid :: "[triple set,etriple ] => bool" ("_ |\<Turnstile>\<^sub>e/ _"[61,61] 60) where
+definition cenvalid :: "[triple set,etriple ] => bool" (\<open>_ |\<Turnstile>\<^sub>e/ _\<close>[61,61] 60) where
"A |\<Turnstile>\<^sub>e t \<equiv> \<forall>n. |\<Turnstile>n: A --> \<Turnstile>n:\<^sub>e t"
lemma nvalid_def2: "\<Turnstile>n: (P,c,Q) \<equiv> \<forall>s t. s -c-n\<rightarrow> t \<longrightarrow> P s \<longrightarrow> Q t"
--- 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 (\<open>Nat\<close>) (* with mixfix because of clash with NatDef.Nat *)
and pred :: fname
and suc add :: mname
and any :: vname
abbreviation
- dummy :: expr ("<>")
+ dummy :: expr (\<open><>\<close>)
where "<> == LAcc any"
abbreviation
@@ -92,7 +92,7 @@
subsection "``atleast'' relation for interpretation of Nat ``values''"
-primrec Nat_atleast :: "state \<Rightarrow> val \<Rightarrow> nat \<Rightarrow> bool" ("_:_ \<ge> _" [51, 51, 51] 50) where
+primrec Nat_atleast :: "state \<Rightarrow> val \<Rightarrow> nat \<Rightarrow> bool" (\<open>_:_ \<ge> _\<close> [51, 51, 51] 50) where
"s:x\<ge>0 = (x\<noteq>Null)"
| "s:x\<ge>Suc n = (\<exists>a. x=Addr a \<and> heap s a \<noteq> None \<and> s:get_field s a pred\<ge>n)"
--- 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" ("_ -_-_\<rightarrow> _" [98,90, 65,98] 89)
- and eval :: "[state,expr,val,nat,state] => bool" ("_ -_\<succ>_-_\<rightarrow> _"[98,95,99,65,98] 89)
+ exec :: "[state,stmt, nat,state] => bool" (\<open>_ -_-_\<rightarrow> _\<close> [98,90, 65,98] 89)
+ and eval :: "[state,expr,val,nat,state] => bool" (\<open>_ -_\<succ>_-_\<rightarrow> _\<close>[98,95,99,65,98] 89)
where
Skip: " s -Skip-n\<rightarrow> s"
--- 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' \<equiv> s' (| locals := locals s |)"
-definition get_local :: "state => vname => val" ("_<_>" [99,0] 99) where
+definition get_local :: "state => vname => val" (\<open>_<_>\<close> [99,0] 99) where
"get_local s x \<equiv> the (locals s x)"
\<comment> \<open>local function:\<close>
@@ -70,10 +70,10 @@
"get_field s a f \<equiv> the (snd (get_obj s a) f)"
\<comment> \<open>local function:\<close>
-definition hupd :: "loc => obj => state => state" ("hupd'(_\<mapsto>_')" [10,10] 1000) where
+definition hupd :: "loc => obj => state => state" (\<open>hupd'(_\<mapsto>_')\<close> [10,10] 1000) where
"hupd a obj s \<equiv> s (| heap := ((heap s)(a\<mapsto>obj))|)"
-definition lupd :: "vname => val => state => state" ("lupd'(_\<mapsto>_')" [10,10] 1000) where
+definition lupd :: "vname => val => state => state" (\<open>lupd'(_\<mapsto>_')\<close> [10,10] 1000) where
"lupd x v s \<equiv> s (| locals := ((locals s)(x\<mapsto>v ))|)"
definition new_obj :: "loc => cname => state => state" where
--- 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 \<comment> \<open>empty statement\<close>
- | 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) \<comment> \<open>local assignment\<close>
- | FAss expr fname expr ("_.._:==_" [95,99,95] 94) \<comment> \<open>field assignment\<close>
+ | Comp stmt stmt (\<open>_;; _\<close> [91,90 ] 90)
+ | Cond expr stmt stmt (\<open>If '(_') _ Else _\<close> [ 3,91,91] 91)
+ | Loop vname stmt (\<open>While '(_') _\<close> [ 3,91 ] 91)
+ | LAss vname expr (\<open>_ :== _\<close> [99, 95] 94) \<comment> \<open>local assignment\<close>
+ | FAss expr fname expr (\<open>_.._:==_\<close> [95,99,95] 94) \<comment> \<open>field assignment\<close>
| Meth "cname \<times> mname" \<comment> \<open>virtual method\<close>
| Impl "cname \<times> mname" \<comment> \<open>method implementation\<close>
and expr
- = NewC cname ("new _" [ 99] 95) \<comment> \<open>object creation\<close>
+ = NewC cname (\<open>new _\<close> [ 99] 95) \<comment> \<open>object creation\<close>
| Cast cname expr \<comment> \<open>type cast\<close>
| LAcc vname \<comment> \<open>local access\<close>
- | FAcc expr fname ("_.._" [95,99] 95) \<comment> \<open>field access\<close>
+ | FAcc expr fname (\<open>_.._\<close> [95,99] 95) \<comment> \<open>field access\<close>
| Call cname expr mname expr
- ("{_}_.._'(_')" [99,95,99,95] 95) \<comment> \<open>method call\<close>
+ (\<open>{_}_.._'(_')\<close> [99,95,99,95] 95) \<comment> \<open>method call\<close>
end
--- 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 \<equiv> {(C,D). C\<noteq>Object \<and> (\<exists>c. class C = Some c \<and> super c=D)}"
abbreviation
- subcls1_syntax :: "[cname, cname] => bool" ("_ \<prec>C1 _" [71,71] 70)
+ subcls1_syntax :: "[cname, cname] => bool" (\<open>_ \<prec>C1 _\<close> [71,71] 70)
where "C \<prec>C1 D == (C,D) \<in> subcls1"
abbreviation
- subcls_syntax :: "[cname, cname] => bool" ("_ \<preceq>C _" [71,71] 70)
+ subcls_syntax :: "[cname, cname] => bool" (\<open>_ \<preceq>C _\<close> [71,71] 70)
where "C \<preceq>C D \<equiv> (C,D) \<in> subcls1\<^sup>*"
@@ -26,7 +26,7 @@
text\<open>Widening, viz. method invocation conversion\<close>
inductive
- widen :: "ty => ty => bool" ("_ \<preceq> _" [71,71] 70)
+ widen :: "ty => ty => bool" (\<open>_ \<preceq> _\<close> [71,71] 70)
where
refl [intro!, simp]: "T \<preceq> T"
| subcls: "C\<preceq>C D \<Longrightarrow> Class C \<preceq> Class D"
--- 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 "\<guillemotleft>name\<guillemotright>lam" ("LAM [_]._")
+| LAM "\<guillemotleft>name\<guillemotright>lam" (\<open>LAM [_]._\<close>)
| NUM "nat"
-| DIFF "lam" "lam" ("_ -- _") (* subtraction *)
-| PLUS "lam" "lam" ("_ ++ _") (* addition *)
+| DIFF "lam" "lam" (\<open>_ -- _\<close>) (* subtraction *)
+| PLUS "lam" "lam" (\<open>_ ++ _\<close>) (* addition *)
| TRUE
| FALSE
| IF "lam" "lam" "lam"
-| FIX "\<guillemotleft>name\<guillemotright>lam" ("FIX [_]._") (* recursion *)
+| FIX "\<guillemotleft>name\<guillemotright>lam" (\<open>FIX [_]._\<close>) (* recursion *)
| ZET "lam" (* zero test *)
| EQI "lam" "lam" (* equality test on numbers *)
section \<open>Capture-Avoiding Substitution\<close>
nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
+ subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" (\<open>_[_::=_]\<close> [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 \<open>Evaluation Contexts\<close>
datatype ctx =
- Hole ("\<box>")
+ Hole (\<open>\<box>\<close>)
| CAPPL "ctx" "lam"
| CAPPR "lam" "ctx"
| CDIFFL "ctx" "lam"
@@ -96,7 +96,7 @@
text \<open>The operation of filling a term into a context:\<close>
fun
- filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>")
+ filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" (\<open>_\<lbrakk>_\<rbrakk>\<close>)
where
"\<box>\<lbrakk>t\<rbrakk> = t"
| "(CAPPL E t')\<lbrakk>t\<rbrakk> = APP (E\<lbrakk>t\<rbrakk>) t'"
@@ -113,7 +113,7 @@
text \<open>The operation of composing two contexts:\<close>
fun
- ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<circ> _")
+ ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" (\<open>_ \<circ> _\<close>)
where
"\<box> \<circ> E' = E'"
| "(CAPPL E t') \<circ> E' = CAPPL (E \<circ> E') t'"
@@ -134,7 +134,7 @@
text \<open>Composing a list (stack) of contexts.\<close>
fun
- ctx_composes :: "ctx list \<Rightarrow> ctx" ("_\<down>")
+ ctx_composes :: "ctx list \<Rightarrow> ctx" (\<open>_\<down>\<close>)
where
"[]\<down> = \<box>"
| "(E#Es)\<down> = (Es\<down>) \<circ> E"
@@ -152,7 +152,7 @@
equivariance val
inductive
- machine :: "lam\<Rightarrow>ctx list\<Rightarrow>lam\<Rightarrow>ctx list\<Rightarrow>bool" ("<_,_> \<mapsto> <_,_>")
+ machine :: "lam\<Rightarrow>ctx list\<Rightarrow>lam\<Rightarrow>ctx list\<Rightarrow>bool" (\<open><_,_> \<mapsto> <_,_>\<close>)
where
m1[intro]: "<APP e1 e2,Es> \<mapsto> <e1,(CAPPL \<box> e2)#Es>"
| m2[intro]: "val v \<Longrightarrow> <v,(CAPPL \<box> e2)#Es> \<mapsto> <e2,(CAPPR v \<box>)#Es>"
@@ -176,7 +176,7 @@
| mH[intro]: "n1\<noteq>n2 \<Longrightarrow> <NUM n1,(CEQIR (NUM n2) \<box>)#Es> \<mapsto> <FALSE,Es>"
inductive
- "machine_star" :: "lam\<Rightarrow>ctx list\<Rightarrow>lam\<Rightarrow>ctx list\<Rightarrow>bool" ("<_,_> \<mapsto>* <_,_>")
+ "machine_star" :: "lam\<Rightarrow>ctx list\<Rightarrow>lam\<Rightarrow>ctx list\<Rightarrow>bool" (\<open><_,_> \<mapsto>* <_,_>\<close>)
where
ms1[intro]: "<e,Es> \<mapsto>* <e,Es>"
| ms2[intro]: "\<lbrakk><e1,Es1> \<mapsto> <e2,Es2>; <e2,Es2> \<mapsto>* <e3,Es3>\<rbrakk> \<Longrightarrow> <e1,Es1> \<mapsto>* <e3,Es3>"
@@ -194,7 +194,7 @@
section \<open>The Evaluation Relation (Big-Step Semantics)\<close>
inductive
- eval :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<Down> _")
+ eval :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open>_ \<Down> _\<close>)
where
eval_NUM[intro]: "NUM n \<Down> NUM n"
| eval_DIFF[intro]: "\<lbrakk>t1 \<Down> (NUM n1); t2 \<Down> (NUM n2)\<rbrakk> \<Longrightarrow> t1 -- t2 \<Down> NUM (n1 - n2)"
@@ -257,7 +257,7 @@
by (simp add: perm_nat_def perm_bool)
inductive
- cbv :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>cbv _")
+ cbv :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open>_ \<longrightarrow>cbv _\<close>)
where
cbv1: "\<lbrakk>val v; x\<sharp>v\<rbrakk> \<Longrightarrow> APP (LAM [x].t) v \<longrightarrow>cbv t[x::=v]"
| cbv2[intro]: "t \<longrightarrow>cbv t' \<Longrightarrow> APP t t2 \<longrightarrow>cbv APP t' t2"
@@ -297,7 +297,7 @@
qed
inductive
- "cbv_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>cbv* _")
+ "cbv_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>cbv* _\<close>)
where
cbvs1[intro]: "e \<longrightarrow>cbv* e"
| cbvs2[intro]: "\<lbrakk>e1\<longrightarrow>cbv e2; e2 \<longrightarrow>cbv* e3\<rbrakk> \<Longrightarrow> e1 \<longrightarrow>cbv* e3"
@@ -366,7 +366,7 @@
tVAR "string"
| tBOOL
| tINT
-| tARR "ty" "ty" ("_ \<rightarrow> _")
+| tARR "ty" "ty" (\<open>_ \<rightarrow> _\<close>)
declare ty.inject[simp]
@@ -384,7 +384,7 @@
text \<open>Sub-Typing Contexts\<close>
abbreviation
- "sub_tctx" :: "tctx \<Rightarrow> tctx \<Rightarrow> bool" ("_ \<subseteq> _")
+ "sub_tctx" :: "tctx \<Rightarrow> tctx \<Rightarrow> bool" (\<open>_ \<subseteq> _\<close>)
where
"\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
@@ -425,7 +425,7 @@
section \<open>The Typing Relation\<close>
inductive
- typing :: "tctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _")
+ typing :: "tctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close>)
where
t_VAR[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> VAR x : T"
| t_APP[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> APP t\<^sub>1 t\<^sub>2 : T\<^sub>2"
--- 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 \<open>Beta Reduction\<close>
inductive
- "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
+ "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>\<beta> _\<close> [80,80] 80)
where
b1[intro]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^sub>\<beta>(App s2 t)"
| b2[intro]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^sub>\<beta>(App t s2)"
@@ -159,7 +159,7 @@
by (simp_all add: abs_fresh fresh_fact')
inductive
- "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _" [80,80] 80)
+ "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _\<close> [80,80] 80)
where
bs1[intro, simp]: "M \<longrightarrow>\<^sub>\<beta>\<^sup>* M"
| bs2[intro]: "\<lbrakk>M1\<longrightarrow>\<^sub>\<beta>\<^sup>* M2; M2 \<longrightarrow>\<^sub>\<beta> M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^sub>\<beta>\<^sup>* M3"
@@ -176,7 +176,7 @@
section \<open>One-Reduction\<close>
inductive
- One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1 _" [80,80] 80)
+ One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>1 _\<close> [80,80] 80)
where
o1[intro!]: "M\<longrightarrow>\<^sub>1M"
| o2[simp,intro!]: "\<lbrakk>t1\<longrightarrow>\<^sub>1t2;s1\<longrightarrow>\<^sub>1s2\<rbrakk> \<Longrightarrow> (App t1 s1)\<longrightarrow>\<^sub>1(App t2 s2)"
@@ -189,7 +189,7 @@
by (simp_all add: abs_fresh fresh_fact')
inductive
- "One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1\<^sup>* _" [80,80] 80)
+ "One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>1\<^sup>* _\<close> [80,80] 80)
where
os1[intro, simp]: "M \<longrightarrow>\<^sub>1\<^sup>* M"
| os2[intro]: "\<lbrakk>M1\<longrightarrow>\<^sub>1\<^sup>* M2; M2 \<longrightarrow>\<^sub>1 M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>\<^sub>1\<^sup>* M3"
--- 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 "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>lam" (\<open>Lam [_]._\<close> [100,100] 100)
nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
+ subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" (\<open>_[_::=_]\<close> [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 \<open>Beta-Reduction\<close>
inductive
- "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
+ "Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>\<beta> _\<close> [80,80] 80)
where
b1[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> App t1 s \<longrightarrow>\<^sub>\<beta> App t2 s"
| b2[intro]: "s1 \<longrightarrow>\<^sub>\<beta> s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^sub>\<beta> App t s2"
@@ -76,7 +76,7 @@
section \<open>Transitive Closure of Beta\<close>
inductive
- "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _" [80,80] 80)
+ "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _\<close> [80,80] 80)
where
bs1[intro]: "t \<longrightarrow>\<^sub>\<beta>\<^sup>* t"
| bs2[intro]: "t \<longrightarrow>\<^sub>\<beta> s \<Longrightarrow> t \<longrightarrow>\<^sub>\<beta>\<^sup>* s"
@@ -85,7 +85,7 @@
section \<open>One-Reduction\<close>
inductive
- One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1 _" [80,80] 80)
+ One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>1 _\<close> [80,80] 80)
where
o1[intro]: "Var x\<longrightarrow>\<^sub>1 Var x"
| o2[intro]: "\<lbrakk>t1\<longrightarrow>\<^sub>1t2; s1\<longrightarrow>\<^sub>1s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>\<^sub>1 App t2 s2"
@@ -148,7 +148,7 @@
section \<open>Transitive Closure of One\<close>
inductive
- "One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1\<^sup>* _" [80,80] 80)
+ "One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>1\<^sup>* _\<close> [80,80] 80)
where
os1[intro]: "t \<longrightarrow>\<^sub>1\<^sup>* t"
| os2[intro]: "t \<longrightarrow>\<^sub>1 s \<Longrightarrow> t \<longrightarrow>\<^sub>1\<^sup>* s"
@@ -157,7 +157,7 @@
section \<open>Complete Development Reduction\<close>
inductive
- Dev :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>\<^sub>d _" [80,80]80)
+ Dev :: "lam \<Rightarrow> lam \<Rightarrow> bool" (\<open> _ \<longrightarrow>\<^sub>d _\<close> [80,80]80)
where
d1[intro]: "Var x \<longrightarrow>\<^sub>d Var x"
| d2[intro]: "t \<longrightarrow>\<^sub>d s \<Longrightarrow> Lam [x].t \<longrightarrow>\<^sub>d Lam[x].s"
--- 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 \<open>types\<close>
-no_notation not ("NOT")
+no_notation not (\<open>NOT\<close>)
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" (\<open>_ AND _\<close> [100,100] 100)
+ | OR "ty" "ty" (\<open>_ OR _\<close> [100,100] 100)
+ | IMP "ty" "ty" (\<open>_ IMP _\<close> [100,100] 100)
instantiation ty :: size
begin
@@ -50,30 +50,30 @@
nominal_datatype trm =
Ax "name" "coname"
- | Cut "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" ("Cut <_>._ ('(_'))._" [0,0,0,100] 101)
- | NotR "\<guillemotleft>name\<guillemotright>trm" "coname" ("NotR ('(_'))._ _" [0,100,100] 101)
- | NotL "\<guillemotleft>coname\<guillemotright>trm" "name" ("NotL <_>._ _" [0,100,100] 101)
- | AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" ("AndR <_>._ <_>._ _" [0,100,0,100,100] 101)
- | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name" ("AndL1 (_)._ _" [100,100,100] 101)
- | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name" ("AndL2 (_)._ _" [100,100,100] 101)
- | OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname" ("OrR1 <_>._ _" [100,100,100] 101)
- | OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname" ("OrR2 <_>._ _" [100,100,100] 101)
- | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" ("OrL (_)._ (_)._ _" [100,100,100,100,100] 101)
- | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname" ("ImpR (_).<_>._ _" [100,100,100,100] 101)
- | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 101)
+ | Cut "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" (\<open>Cut <_>._ ('(_'))._\<close> [0,0,0,100] 101)
+ | NotR "\<guillemotleft>name\<guillemotright>trm" "coname" (\<open>NotR ('(_'))._ _\<close> [0,100,100] 101)
+ | NotL "\<guillemotleft>coname\<guillemotright>trm" "name" (\<open>NotL <_>._ _\<close> [0,100,100] 101)
+ | AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" (\<open>AndR <_>._ <_>._ _\<close> [0,100,0,100,100] 101)
+ | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name" (\<open>AndL1 (_)._ _\<close> [100,100,100] 101)
+ | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name" (\<open>AndL2 (_)._ _\<close> [100,100,100] 101)
+ | OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname" (\<open>OrR1 <_>._ _\<close> [100,100,100] 101)
+ | OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname" (\<open>OrR2 <_>._ _\<close> [100,100,100] 101)
+ | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" (\<open>OrL (_)._ (_)._ _\<close> [100,100,100,100,100] 101)
+ | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname" (\<open>ImpR (_).<_>._ _\<close> [100,100,100,100] 101)
+ | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" (\<open>ImpL <_>._ (_)._ _\<close> [100,100,100,100,100] 101)
text \<open>named terms\<close>
-nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" ("((_):_)" [100,100] 100)
+nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" (\<open>((_):_)\<close> [100,100] 100)
text \<open>conamed terms\<close>
-nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" ("(<_>:_)" [100,100] 100)
+nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" (\<open>(<_>:_)\<close> [100,100] 100)
text \<open>renaming functions\<close>
nominal_primrec (freshness_context: "(d::coname,e::coname)")
- crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm" ("_[_\<turnstile>c>_]" [100,0,0] 100)
+ crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm" (\<open>_[_\<turnstile>c>_]\<close> [100,0,0] 100)
where
"(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)"
| "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>M\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N)[d\<turnstile>c>e] = Cut <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>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 \<Rightarrow> name \<Rightarrow> name \<Rightarrow> trm" ("_[_\<turnstile>n>_]" [100,0,0] 100)
+ nrename :: "trm \<Rightarrow> name \<Rightarrow> name \<Rightarrow> trm" (\<open>_[_\<turnstile>n>_]\<close> [100,0,0] 100)
where
"(Ax x a)[u\<turnstile>n>v] = (if x=u then Ax v a else Ax x a)"
| "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])"
@@ -610,7 +610,7 @@
qed
nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)")
- substn :: "trm \<Rightarrow> name \<Rightarrow> coname \<Rightarrow> trm \<Rightarrow> trm" ("_{_:=<_>._}" [100,100,100,100] 100)
+ substn :: "trm \<Rightarrow> name \<Rightarrow> coname \<Rightarrow> trm \<Rightarrow> trm" (\<open>_{_:=<_>._}\<close> [100,100,100,100] 100)
where
"(Ax x a){y:=<c>.P} = (if x=y then Cut <c>.P (y).Ax y a else Ax x a)"
| "\<lbrakk>a\<sharp>(c,P,N);x\<sharp>(y,P,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N){y:=<c>.P} =
@@ -673,7 +673,7 @@
done
nominal_primrec (freshness_context: "(d::name,z::coname,P::trm)")
- substc :: "trm \<Rightarrow> coname \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_{_:=('(_'))._}" [100,0,0,100] 100)
+ substc :: "trm \<Rightarrow> coname \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" (\<open>_{_:=('(_'))._}\<close> [100,0,0,100] 100)
where
"(Ax x a){d:=(z).P} = (if d=a then Cut <a>.(Ax x a) (z).P else Ax x a)"
| "\<lbrakk>a\<sharp>(d,P,N);x\<sharp>(z,P,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N){d:=(z).P} =
@@ -2423,7 +2423,7 @@
qed (use fic_rest_elims in force)+
inductive
- l_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>l _" [100,100] 100)
+ l_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longrightarrow>\<^sub>l _\<close> [100,100] 100)
where
LAxR: "\<lbrakk>x\<sharp>M; a\<sharp>b; fic M a\<rbrakk> \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^sub>l M[a\<turnstile>c>b]"
| LAxL: "\<lbrakk>a\<sharp>M; x\<sharp>y; fin M x\<rbrakk> \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^sub>l M[x\<turnstile>n>y]"
@@ -3044,7 +3044,7 @@
inductive
- c_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>c _" [100,100] 100)
+ c_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longrightarrow>\<^sub>c _\<close> [100,100] 100)
where
left[intro]: "\<lbrakk>\<not>fic M a; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c M{a:=(x).N}"
| right[intro]: "\<lbrakk>\<not>fin N x; a\<sharp>N; x\<sharp>M\<rbrakk> \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^sub>c N{x:=<a>.M}"
@@ -3098,7 +3098,7 @@
by (induct rule: c_redu.induct) (auto simp: abs_fresh rename_fresh subst_fresh)+
inductive
- a_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>a _" [100,100] 100)
+ a_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longrightarrow>\<^sub>a _\<close> [100,100] 100)
where
al_redu[intro]: "M\<longrightarrow>\<^sub>l M' \<Longrightarrow> M \<longrightarrow>\<^sub>a M'"
| ac_redu[intro]: "M\<longrightarrow>\<^sub>c M' \<Longrightarrow> M \<longrightarrow>\<^sub>a M'"
@@ -3477,7 +3477,7 @@
text \<open>Transitive Closure\<close>
abbreviation
- a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^sub>a* _" [80,80] 80)
+ a_star_redu :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longrightarrow>\<^sub>a* _\<close> [80,80] 80)
where
"M \<longrightarrow>\<^sub>a* M' \<equiv> (a_redu)\<^sup>*\<^sup>* M M'"
--- 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 \<Rightarrow> ntrm set" ("\<parallel>'(_')\<parallel>" [60] 60)
+ CANDn::"ty \<Rightarrow> ntrm set" (\<open>\<parallel>'(_')\<parallel>\<close> [60] 60)
where
"\<parallel>(B)\<parallel> \<equiv> lfp (NEGn B \<circ> NEGc B)"
abbreviation
- CANDc::"ty \<Rightarrow> ctrm set" ("\<parallel><_>\<parallel>" [60] 60)
+ CANDc::"ty \<Rightarrow> ctrm set" (\<open>\<parallel><_>\<parallel>\<close> [60] 60)
where
"\<parallel><B>\<parallel> \<equiv> NEGc B (\<parallel>(B)\<parallel>)"
--- 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\<times>coname\<times>trm) list \<Rightarrow> name \<Rightarrow> (coname\<times>trm) option \<Rightarrow> bool" ("_ nmaps _ to _" [55,55,55] 55)
+ nmaps :: "(name\<times>coname\<times>trm) list \<Rightarrow> name \<Rightarrow> (coname\<times>trm) option \<Rightarrow> bool" (\<open>_ nmaps _ to _\<close> [55,55,55] 55)
where
"\<theta>_n nmaps x to P \<equiv> (findn \<theta>_n x) = P"
abbreviation
- cmaps :: "(coname\<times>name\<times>trm) list \<Rightarrow> coname \<Rightarrow> (name\<times>trm) option \<Rightarrow> bool" ("_ cmaps _ to _" [55,55,55] 55)
+ cmaps :: "(coname\<times>name\<times>trm) list \<Rightarrow> coname \<Rightarrow> (name\<times>trm) option \<Rightarrow> bool" (\<open>_ cmaps _ to _\<close> [55,55,55] 55)
where
"\<theta>_c cmaps a to P \<equiv> (findc \<theta>_c a) = P"
@@ -3662,7 +3662,7 @@
done
nominal_primrec (freshness_context: "(\<theta>_n::(name\<times>coname\<times>trm) list,\<theta>_c::(coname\<times>name\<times>trm) list)")
- psubst :: "(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm\<Rightarrow>trm" ("_,_<_>" [100,100,100] 100)
+ psubst :: "(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm\<Rightarrow>trm" (\<open>_,_<_>\<close> [100,100,100] 100)
where
"\<theta>_n,\<theta>_c<Ax x a> = lookup x a \<theta>_n \<theta>_c"
| "\<lbrakk>a\<sharp>(N,\<theta>_n,\<theta>_c);x\<sharp>(M,\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> \<theta>_n,\<theta>_c<Cut <a>.M (x).N> =
@@ -4863,12 +4863,12 @@
done
definition
- ncloses :: "(name\<times>coname\<times>trm) list\<Rightarrow>(name\<times>ty) list \<Rightarrow> bool" ("_ ncloses _" [55,55] 55)
+ ncloses :: "(name\<times>coname\<times>trm) list\<Rightarrow>(name\<times>ty) list \<Rightarrow> bool" (\<open>_ ncloses _\<close> [55,55] 55)
where
"\<theta>_n ncloses \<Gamma> \<equiv> \<forall>x B. ((x,B) \<in> set \<Gamma> \<longrightarrow> (\<exists>c P. \<theta>_n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)))"
definition
- ccloses :: "(coname\<times>name\<times>trm) list\<Rightarrow>(coname\<times>ty) list \<Rightarrow> bool" ("_ ccloses _" [55,55] 55)
+ ccloses :: "(coname\<times>name\<times>trm) list\<Rightarrow>(coname\<times>ty) list \<Rightarrow> bool" (\<open>_ ccloses _\<close> [55,55] 55)
where
"\<theta>_c ccloses \<Delta> \<equiv> \<forall>a B. ((a,B) \<in> set \<Delta> \<longrightarrow> (\<exists>x P. \<theta>_c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)))"
@@ -4938,7 +4938,7 @@
text \<open>typing relation\<close>
inductive
- typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" ("_ \<turnstile> _ \<turnstile> _" [100,100,100] 100)
+ typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" (\<open>_ \<turnstile> _ \<turnstile> _\<close> [100,100,100] 100)
where
TAx: "\<lbrakk>validn \<Gamma>;validc \<Delta>; (x,B)\<in>set \<Gamma>; (a,B)\<in>set \<Delta>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Ax x a \<turnstile> \<Delta>"
| TNotR: "\<lbrakk>x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Delta>' = {(a,NOT B)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk>
--- 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" ("_\<rightarrow>_" [100,100] 100)
+ | Arrow "ty" "ty" (\<open>_\<rightarrow>_\<close> [100,100] 100)
nominal_datatype trm =
Var "name"
- | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>trm" (\<open>Lam [_]._\<close> [100,100] 100)
| App "trm" "trm"
| Const "nat"
| Pr "trm" "trm"
@@ -26,24 +26,24 @@
| InL "trm"
| InR "trm"
| Case "trm" "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"
- ("Case _ of inl _ \<rightarrow> _ | inr _ \<rightarrow> _" [100,100,100,100,100] 100)
+ (\<open>Case _ of inl _ \<rightarrow> _ | inr _ \<rightarrow> _\<close> [100,100,100,100,100] 100)
nominal_datatype dataI = OneI | NatI
nominal_datatype tyI =
DataI "dataI"
- | ArrowI "tyI" "tyI" ("_\<rightarrow>_" [100,100] 100)
+ | ArrowI "tyI" "tyI" (\<open>_\<rightarrow>_\<close> [100,100] 100)
nominal_datatype trmI =
IVar "name"
- | ILam "\<guillemotleft>name\<guillemotright>trmI" ("ILam [_]._" [100,100] 100)
+ | ILam "\<guillemotleft>name\<guillemotright>trmI" (\<open>ILam [_]._\<close> [100,100] 100)
| IApp "trmI" "trmI"
| IUnit
| INat "nat"
| ISucc "trmI"
- | IAss "trmI" "trmI" ("_\<mapsto>_" [100,100] 100)
+ | IAss "trmI" "trmI" (\<open>_\<mapsto>_\<close> [100,100] 100)
| IRef "trmI"
- | ISeq "trmI" "trmI" ("_;;_" [100,100] 100)
+ | ISeq "trmI" "trmI" (\<open>_;;_\<close> [100,100] 100)
| Iif "trmI" "trmI" "trmI"
text \<open>valid contexts\<close>
@@ -57,7 +57,7 @@
text \<open>typing judgements for trms\<close>
inductive
- typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80)
+ typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (\<open> _ \<turnstile> _ : _ \<close> [80,80,80] 80)
where
t0[intro]: "\<lbrakk>valid \<Gamma>; (x,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : \<tau>"
| t1[intro]: "\<lbrakk>\<Gamma> \<turnstile> e1 : \<tau>1\<rightarrow>\<tau>2; \<Gamma> \<turnstile> e2 : \<tau>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App e1 e2 : \<tau>2"
@@ -75,7 +75,7 @@
text \<open>typing judgements for Itrms\<close>
inductive
- Ityping :: "(name\<times>tyI) list\<Rightarrow>trmI\<Rightarrow>tyI\<Rightarrow>bool" (" _ I\<turnstile> _ : _ " [80,80,80] 80)
+ Ityping :: "(name\<times>tyI) list\<Rightarrow>trmI\<Rightarrow>tyI\<Rightarrow>bool" (\<open> _ I\<turnstile> _ : _ \<close> [80,80,80] 80)
where
t0[intro]: "\<lbrakk>valid \<Gamma>; (x,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> I\<turnstile> IVar x : \<tau>"
| t1[intro]: "\<lbrakk>\<Gamma> I\<turnstile> e1 : \<tau>1\<rightarrow>\<tau>2; \<Gamma> I\<turnstile> e2 : \<tau>1\<rbrakk>\<Longrightarrow> \<Gamma> I\<turnstile> IApp e1 e2 : \<tau>2"
@@ -91,7 +91,7 @@
text \<open>capture-avoiding substitution\<close>
class subst =
- fixes subst :: "'a \<Rightarrow> name \<Rightarrow> 'a \<Rightarrow> 'a" ("_[_::=_]" [100,100,100] 100)
+ fixes subst :: "'a \<Rightarrow> name \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>_[_::=_]\<close> [100,100,100] 100)
instantiation trm :: subst
begin
@@ -169,7 +169,7 @@
text \<open>big-step evaluation for trms\<close>
inductive
- big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80)
+ big :: "trm\<Rightarrow>trm\<Rightarrow>bool" (\<open>_ \<Down> _\<close> [80,80] 80)
where
b0[intro]: "Lam [x].e \<Down> Lam [x].e"
| b1[intro]: "\<lbrakk>e1\<Down>Lam [x].e; e2\<Down>e2'; e[x::=e2']\<Down>e'\<rbrakk> \<Longrightarrow> App e1 e2 \<Down> e'"
@@ -183,7 +183,7 @@
| b9[intro]: "\<lbrakk>e\<Down>InR e'; e2[x::=e']\<Down>e''\<rbrakk> \<Longrightarrow> Case e of inl x1 \<rightarrow> e1 | inr x2 \<rightarrow> e2 \<Down> e''"
inductive
- Ibig :: "((nat\<Rightarrow>nat)\<times>trmI)\<Rightarrow>((nat\<Rightarrow>nat)\<times>trmI)\<Rightarrow>bool" ("_ I\<Down> _" [80,80] 80)
+ Ibig :: "((nat\<Rightarrow>nat)\<times>trmI)\<Rightarrow>((nat\<Rightarrow>nat)\<times>trmI)\<Rightarrow>bool" (\<open>_ I\<Down> _\<close> [80,80] 80)
where
m0[intro]: "(m,ILam [x].e) I\<Down> (m,ILam [x].e)"
| m1[intro]: "\<lbrakk>(m,e1)I\<Down>(m',ILam [x].e); (m',e2)I\<Down>(m'',e3); (m'',e[x::=e3])I\<Down>(m''',e4)\<rbrakk>
--- 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 "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>lam" (\<open>Lam [_]._\<close> [100,100] 100)
text \<open>
Contexts - the lambda case in contexts does not bind
@@ -33,15 +33,15 @@
\<close>
nominal_datatype ctx =
- Hole ("\<box>" 1000)
+ Hole (\<open>\<box>\<close> 1000)
| CAppL "ctx" "lam"
| CAppR "lam" "ctx"
- | CLam "name" "ctx" ("CLam [_]._" [100,100] 100)
+ | CLam "name" "ctx" (\<open>CLam [_]._\<close> [100,100] 100)
text \<open>Capture-Avoiding Substitution\<close>
nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
+ subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" (\<open>_[_::=_]\<close> [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 @@
\<close>
nominal_primrec
- filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>" [100,100] 100)
+ filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" (\<open>_\<lbrakk>_\<rbrakk>\<close> [100,100] 100)
where
"\<box>\<lbrakk>t\<rbrakk> = t"
| "(CAppL E t')\<lbrakk>t\<rbrakk> = App (E\<lbrakk>t\<rbrakk>) t'"
@@ -79,7 +79,7 @@
text \<open>The composition of two contexts.\<close>
nominal_primrec
- ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<circ> _" [100,100] 100)
+ ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" (\<open>_ \<circ> _\<close> [100,100] 100)
where
"\<box> \<circ> E' = E'"
| "(CAppL E t') \<circ> E' = CAppL (E \<circ> E') t'"
@@ -94,14 +94,14 @@
text \<open>Beta-reduction via contexts in the Felleisen-Hieb style.\<close>
inductive
- ctx_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>x _" [80,80] 80)
+ ctx_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open>_ \<longrightarrow>x _\<close> [80,80] 80)
where
xbeta[intro]: "E\<lbrakk>App (Lam [x].t) t'\<rbrakk> \<longrightarrow>x E\<lbrakk>t[x::=t']\<rbrakk>"
text \<open>Beta-reduction via congruence rules in the Plotkin style.\<close>
inductive
- cong_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>o _" [80,80] 80)
+ cong_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open>_ \<longrightarrow>o _\<close> [80,80] 80)
where
obeta[intro]: "App (Lam [x].t) t' \<longrightarrow>o t[x::=t']"
| oapp1[intro]: "t \<longrightarrow>o t' \<Longrightarrow> App t t2 \<longrightarrow>o App t' t2"
--- 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" ("_\<rightarrow>_" [100,100] 100)
+ | Arrow "ty" "ty" (\<open>_\<rightarrow>_\<close> [100,100] 100)
nominal_datatype trm =
Unit
- | Var "name" ("Var _" [100] 100)
- | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
- | App "trm" "trm" ("App _ _" [110,110] 100)
+ | Var "name" (\<open>Var _\<close> [100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>trm" (\<open>Lam [_]._\<close> [100,100] 100)
+ | App "trm" "trm" (\<open>App _ _\<close> [110,110] 100)
| Const "nat"
type_synonym Ctxt = "(name\<times>ty) list"
@@ -94,7 +94,7 @@
(auto simp: fresh_list_cons fresh_prod fresh_atm)
nominal_primrec
- psubst :: "Subst \<Rightarrow> trm \<Rightarrow> trm" ("_<_>" [100,100] 130)
+ psubst :: "Subst \<Rightarrow> trm \<Rightarrow> trm" (\<open>_<_>\<close> [100,100] 130)
where
"\<theta><(Var x)> = (lookup \<theta> x)"
| "\<theta><(App t\<^sub>1 t\<^sub>2)> = App \<theta><t\<^sub>1> \<theta><t\<^sub>2>"
@@ -104,7 +104,7 @@
by(finite_guess | simp add: abs_fresh | fresh_guess)+
abbreviation
- subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
+ subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" (\<open>_[_::=_]\<close> [100,100,100] 100)
where
"t[x::=t'] \<equiv> ([(x,t')])<t>"
@@ -268,7 +268,7 @@
valid_cons_elim_auto[elim]:"valid ((x,T)#\<Gamma>)"
abbreviation
- "sub_context" :: "Ctxt \<Rightarrow> Ctxt \<Rightarrow> bool" (" _ \<subseteq> _ " [55,55] 55)
+ "sub_context" :: "Ctxt \<Rightarrow> Ctxt \<Rightarrow> bool" (\<open> _ \<subseteq> _ \<close> [55,55] 55)
where
"\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>a T. (a,T)\<in>set \<Gamma>\<^sub>1 \<longrightarrow> (a,T)\<in>set \<Gamma>\<^sub>2"
@@ -298,7 +298,7 @@
(auto dest!: fresh_context)
inductive
- typing :: "Ctxt\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [60,60,60] 60)
+ typing :: "Ctxt\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (\<open> _ \<turnstile> _ : _ \<close> [60,60,60] 60)
where
T_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
| T_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> e\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> e\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App e\<^sub>1 e\<^sub>2 : T\<^sub>2"
@@ -334,7 +334,7 @@
section \<open>Definitional Equivalence\<close>
inductive
- def_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<equiv> _ : _" [60,60] 60)
+ def_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (\<open>_ \<turnstile> _ \<equiv> _ : _\<close> [60,60] 60)
where
Q_Refl[intro]: "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<equiv> t : T"
| Q_Symm[intro]: "\<Gamma> \<turnstile> t \<equiv> s : T \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"
@@ -360,7 +360,7 @@
section \<open>Weak Head Reduction\<close>
inductive
- whr_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<leadsto> _" [80,80] 80)
+ whr_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" (\<open>_ \<leadsto> _\<close> [80,80] 80)
where
QAR_Beta[intro]: "App (Lam [x]. t\<^sub>1) t\<^sub>2 \<leadsto> t\<^sub>1[x::=t\<^sub>2]"
| QAR_App[intro]: "t\<^sub>1 \<leadsto> t\<^sub>1' \<Longrightarrow> App t\<^sub>1 t\<^sub>2 \<leadsto> App t\<^sub>1' t\<^sub>2"
@@ -387,12 +387,12 @@
section \<open>Weak Head Normalisation\<close>
abbreviation
- nf :: "trm \<Rightarrow> bool" ("_ \<leadsto>|" [100] 100)
+ nf :: "trm \<Rightarrow> bool" (\<open>_ \<leadsto>|\<close> [100] 100)
where
"t\<leadsto>| \<equiv> \<not>(\<exists> u. t \<leadsto> u)"
inductive
- whn_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80)
+ whn_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" (\<open>_ \<Down> _\<close> [80,80] 80)
where
QAN_Reduce[intro]: "\<lbrakk>s \<leadsto> t; t \<Down> u\<rbrakk> \<Longrightarrow> s \<Down> u"
| QAN_Normal[intro]: "t\<leadsto>| \<Longrightarrow> t \<Down> t"
@@ -429,9 +429,9 @@
section \<open>Algorithmic Term Equivalence and Algorithmic Path Equivalence\<close>
inductive
- alg_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<Leftrightarrow> _ : _" [60,60,60,60] 60)
+ alg_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (\<open>_ \<turnstile> _ \<Leftrightarrow> _ : _\<close> [60,60,60,60] 60)
and
- alg_path_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<leftrightarrow> _ : _" [60,60,60,60] 60)
+ alg_path_equiv :: "Ctxt\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (\<open>_ \<turnstile> _ \<leftrightarrow> _ : _\<close> [60,60,60,60] 60)
where
QAT_Base[intro]: "\<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase"
| QAT_Arrow[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T\<^sub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^sub>2\<rbrakk>
@@ -592,7 +592,7 @@
section \<open>Logical Equivalence\<close>
-function log_equiv :: "(Ctxt \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool)" ("_ \<turnstile> _ is _ : _" [60,60,60,60] 60)
+function log_equiv :: "(Ctxt \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool)" (\<open>_ \<turnstile> _ is _ : _\<close> [60,60,60,60] 60)
where
"\<Gamma> \<turnstile> s is t : TUnit = True"
| "\<Gamma> \<turnstile> s is t : TBase = \<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase"
@@ -744,7 +744,7 @@
qed
abbreviation
- log_equiv_for_psubsts :: "Ctxt \<Rightarrow> Subst \<Rightarrow> Subst \<Rightarrow> Ctxt \<Rightarrow> bool" ("_ \<turnstile> _ is _ over _" [60,60] 60)
+ log_equiv_for_psubsts :: "Ctxt \<Rightarrow> Subst \<Rightarrow> Subst \<Rightarrow> Ctxt \<Rightarrow> bool" (\<open>_ \<turnstile> _ is _ over _\<close> [60,60] 60)
where
"\<Gamma>' \<turnstile> \<theta> is \<theta>' over \<Gamma> \<equiv> \<forall>x T. (x,T) \<in> set \<Gamma> \<longrightarrow> \<Gamma>' \<turnstile> \<theta><Var x> is \<theta>'<Var x> : T"
--- 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 \<open>Types for Names, Nominal Datatype Declaration for Types and Terms\<close>
no_syntax
- "_Map" :: "maplets => 'a \<rightharpoonup> 'b" ("(1[_])")
+ "_Map" :: "maplets => 'a \<rightharpoonup> 'b" (\<open>(1[_])\<close>)
text \<open>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 "\<rightarrow>" 200)
+ | Arrow "ty" "ty" (infixr \<open>\<rightarrow>\<close> 200)
| Forall "\<guillemotleft>tyvrs\<guillemotright>ty" "ty"
nominal_datatype trm =
Var "vrs"
| Abs "\<guillemotleft>vrs\<guillemotright>trm" "ty"
| TAbs "\<guillemotleft>tyvrs\<guillemotright>trm" "ty"
- | App "trm" "trm" (infixl "\<cdot>" 200)
- | TApp "trm" "ty" (infixl "\<cdot>\<^sub>\<tau>" 200)
+ | App "trm" "trm" (infixl \<open>\<cdot>\<close> 200)
+ | TApp "trm" "ty" (infixl \<open>\<cdot>\<^sub>\<tau>\<close> 200)
text \<open>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>\<open>Arrow\<close>.\<close>
abbreviation
- Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty" ("(3\<forall>_<:_./ _)" [0, 0, 10] 10)
+ Forall_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty" (\<open>(3\<forall>_<:_./ _)\<close> [0, 0, 10] 10)
where
"\<forall>X<:T\<^sub>1. T\<^sub>2 \<equiv> ty.Forall X T\<^sub>2 T\<^sub>1"
abbreviation
- Abs_syn :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("(3\<lambda>_:_./ _)" [0, 0, 10] 10)
+ Abs_syn :: "vrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" (\<open>(3\<lambda>_:_./ _)\<close> [0, 0, 10] 10)
where
"\<lambda>x:T. t \<equiv> trm.Abs x t T"
abbreviation
- TAbs_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("(3\<lambda>_<:_./ _)" [0, 0, 10] 10)
+ TAbs_syn :: "tyvrs \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" (\<open>(3\<lambda>_<:_./ _)\<close> [0, 0, 10] 10)
where
"\<lambda>X<:T. t \<equiv> trm.TAbs X t T"
@@ -223,7 +223,7 @@
in \<^term>\<open>\<Gamma>\<close>. The set of free variables of \<^term>\<open>S\<close> is the
\<open>support\<close> of \<^term>\<open>S\<close>.\<close>
-definition "closed_in" :: "ty \<Rightarrow> env \<Rightarrow> bool" ("_ closed'_in _" [100,100] 100) where
+definition "closed_in" :: "ty \<Rightarrow> env \<Rightarrow> bool" (\<open>_ closed'_in _\<close> [100,100] 100) where
"S closed_in \<Gamma> \<equiv> (supp S)\<subseteq>(ty_dom \<Gamma>)"
lemma closed_in_eqvt[eqvt]:
@@ -279,7 +279,7 @@
text \<open>Now validity of a context is a straightforward inductive definition.\<close>
inductive
- valid_rel :: "env \<Rightarrow> bool" ("\<turnstile> _ ok" [100] 100)
+ valid_rel :: "env \<Rightarrow> bool" (\<open>\<turnstile> _ ok\<close> [100] 100)
where
valid_nil[simp]: "\<turnstile> [] ok"
| valid_consT[simp]: "\<lbrakk>\<turnstile> \<Gamma> ok; X\<sharp>(ty_dom \<Gamma>); T closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<turnstile> (TVarB X T#\<Gamma>) ok"
@@ -386,7 +386,7 @@
by (finite_guess | fresh_guess | simp)+
nominal_primrec
- subst_ty :: "ty \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty" ("_[_ \<mapsto> _]\<^sub>\<tau>" [300, 0, 0] 300)
+ subst_ty :: "ty \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty" (\<open>_[_ \<mapsto> _]\<^sub>\<tau>\<close> [300, 0, 0] 300)
where
"(Tvar X)[Y \<mapsto> T]\<^sub>\<tau> = (if X=Y then T else Tvar X)"
| "(Top)[Y \<mapsto> T]\<^sub>\<tau> = Top"
@@ -439,7 +439,7 @@
(simp_all add: fresh_atm calc_atm abs_fresh fresh_aux)
nominal_primrec
- subst_tyb :: "binding \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> binding" ("_[_ \<mapsto> _]\<^sub>b" [100,100,100] 100)
+ subst_tyb :: "binding \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> binding" (\<open>_[_ \<mapsto> _]\<^sub>b\<close> [100,100,100] 100)
where
"(TVarB X U)[Y \<mapsto> T]\<^sub>b = TVarB X (U[Y \<mapsto> T]\<^sub>\<tau>)"
| "(VarB X U)[Y \<mapsto> T]\<^sub>b = VarB X (U[Y \<mapsto> T]\<^sub>\<tau>)"
@@ -459,7 +459,7 @@
by (induct B rule: binding.induct)
(simp_all add: fresh_atm type_subst_identity)
-primrec subst_tyc :: "env \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> env" ("_[_ \<mapsto> _]\<^sub>e" [100,100,100] 100) where
+primrec subst_tyc :: "env \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> env" (\<open>_[_ \<mapsto> _]\<^sub>e\<close> [100,100,100] 100) where
"([])[Y \<mapsto> T]\<^sub>e= []"
| "(B#\<Gamma>)[Y \<mapsto> T]\<^sub>e = (B[Y \<mapsto> T]\<^sub>b)#(\<Gamma>[Y \<mapsto> T]\<^sub>e)"
@@ -485,7 +485,7 @@
by (induct \<Delta>) simp_all
nominal_primrec
- subst_trm :: "trm \<Rightarrow> vrs \<Rightarrow> trm \<Rightarrow> trm" ("_[_ \<mapsto> _]" [300, 0, 0] 300)
+ subst_trm :: "trm \<Rightarrow> vrs \<Rightarrow> trm \<Rightarrow> trm" (\<open>_[_ \<mapsto> _]\<close> [300, 0, 0] 300)
where
"(Var x)[y \<mapsto> t'] = (if x=y then t' else (Var x))"
| "(t1 \<cdot> t2)[y \<mapsto> t'] = t1[y \<mapsto> t'] \<cdot> t2[y \<mapsto> 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 \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> trm" ("_[_ \<mapsto>\<^sub>\<tau> _]" [300, 0, 0] 300)
+ subst_trm_ty :: "trm \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> trm" (\<open>_[_ \<mapsto>\<^sub>\<tau> _]\<close> [300, 0, 0] 300)
where
"(Var x)[Y \<mapsto>\<^sub>\<tau> T2] = Var x"
| "(t1 \<cdot> t2)[Y \<mapsto>\<^sub>\<tau> T2] = t1[Y \<mapsto>\<^sub>\<tau> T2] \<cdot> t2[Y \<mapsto>\<^sub>\<tau> T2]"
@@ -576,7 +576,7 @@
$\alpha$-equivalence classes.)\<close>
inductive
- subtype_of :: "env \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_\<turnstile>_<:_" [100,100,100] 100)
+ subtype_of :: "env \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" (\<open>_\<turnstile>_<:_\<close> [100,100,100] 100)
where
SA_Top[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; S closed_in \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"
| SA_refl_TVar[intro]: "\<lbrakk>\<turnstile> \<Gamma> ok; X \<in> ty_dom \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> 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.\<close>
-definition extends :: "env \<Rightarrow> env \<Rightarrow> bool" ("_ extends _" [100,100] 100) where
+definition extends :: "env \<Rightarrow> env \<Rightarrow> bool" (\<open>_ extends _\<close> [100,100] 100) where
"\<Delta> extends \<Gamma> \<equiv> \<forall>X Q. (TVarB X Q)\<in>set \<Gamma> \<longrightarrow> (TVarB X Q)\<in>set \<Delta>"
lemma extends_ty_dom:
@@ -1017,7 +1017,7 @@
section \<open>Typing\<close>
inductive
- typing :: "env \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
+ typing :: "env \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [60,60,60] 60)
where
T_Var[intro]: "\<lbrakk> VarB x T \<in> set \<Gamma>; \<turnstile> \<Gamma> ok \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
| T_App[intro]: "\<lbrakk> \<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1 \<rightarrow> T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 \<cdot> t\<^sub>2 : T\<^sub>2"
@@ -1145,7 +1145,7 @@
"val (t1 \<cdot>\<^sub>\<tau> t2)"
inductive
- eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longmapsto> _" [60,60] 60)
+ eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longmapsto> _\<close> [60,60] 60)
where
E_Abs : "\<lbrakk> x \<sharp> v\<^sub>2; val v\<^sub>2 \<rbrakk> \<Longrightarrow> (\<lambda>x:T\<^sub>11. t\<^sub>12) \<cdot> v\<^sub>2 \<longmapsto> t\<^sub>12[x \<mapsto> v\<^sub>2]"
| E_App1 [intro]: "t \<longmapsto> t' \<Longrightarrow> t \<cdot> u \<longmapsto> t' \<cdot> u"
--- 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 "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>lam" (\<open>Lam [_]._\<close> [100,100] 100)
text \<open>Definition of the height-function on lambda-terms.\<close>
@@ -31,7 +31,7 @@
text \<open>Definition of capture-avoiding substitution.\<close>
nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
+ subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" (\<open>_[_::=_]\<close> [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'])"
--- 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 "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>lam" (\<open>Lam [_]._\<close> [100,100] 100)
text \<open>The depth of a lambda-term.\<close>
@@ -68,7 +68,7 @@
by (induct \<theta>) (auto simp add: eqvts)
nominal_primrec
- psubst :: "(name\<times>lam) list \<Rightarrow> lam \<Rightarrow> lam" ("_<_>" [95,95] 105)
+ psubst :: "(name\<times>lam) list \<Rightarrow> lam \<Rightarrow> lam" (\<open>_<_>\<close> [95,95] 105)
where
"\<theta><(Var x)> = (lookup \<theta> x)"
| "\<theta><(App e\<^sub>1 e\<^sub>2)> = App (\<theta><e\<^sub>1>) (\<theta><e\<^sub>2>)"
@@ -83,7 +83,7 @@
(simp_all add: eqvts fresh_bij)
abbreviation
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
+ subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" (\<open>_[_::=_]\<close> [100,100,100] 100)
where
"t[x::=t'] \<equiv> ([(x,t')])<t>"
@@ -115,15 +115,15 @@
name, even if we introduce the notation [_]._ for CLam.
\<close>
nominal_datatype clam =
- Hole ("\<box>" 1000)
+ Hole (\<open>\<box>\<close> 1000)
| CAppL "clam" "lam"
| CAppR "lam" "clam"
- | CLam "name" "clam" ("CLam [_]._" [100,100] 100)
+ | CLam "name" "clam" (\<open>CLam [_]._\<close> [100,100] 100)
text \<open>Filling a lambda-term into a context.\<close>
nominal_primrec
- filling :: "clam \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>" [100,100] 100)
+ filling :: "clam \<Rightarrow> lam \<Rightarrow> lam" (\<open>_\<lbrakk>_\<rbrakk>\<close> [100,100] 100)
where
"\<box>\<lbrakk>t\<rbrakk> = t"
| "(CAppL E t')\<lbrakk>t\<rbrakk> = App (E\<lbrakk>t\<rbrakk>) t'"
@@ -134,7 +134,7 @@
text \<open>Composition od two contexts\<close>
nominal_primrec
- clam_compose :: "clam \<Rightarrow> clam \<Rightarrow> clam" ("_ \<circ> _" [100,100] 100)
+ clam_compose :: "clam \<Rightarrow> clam \<Rightarrow> clam" (\<open>_ \<circ> _\<close> [100,100] 100)
where
"\<box> \<circ> E' = E'"
| "(CAppL E t') \<circ> E' = CAppL (E \<circ> E') t'"
--- 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 "\<guillemotleft>var\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>var\<guillemotright>trm" (\<open>Lam [_]._\<close> [100,100] 100)
| App "trm" "trm"
| Pss "mvar" "trm" (* passivate *)
- | Act "\<guillemotleft>mvar\<guillemotright>trm" ("Act [_]._" [100,100] 100) (* activate *)
+ | Act "\<guillemotleft>mvar\<guillemotright>trm" (\<open>Act [_]._\<close> [100,100] 100) (* activate *)
end
--- 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 "\<rightarrow>" 200)
+ | TArr "ty" "ty" (infix \<open>\<rightarrow>\<close> 200)
lemma ty_fresh[simp]:
fixes x::"name"
@@ -99,7 +99,7 @@
text \<open>"weak" typing relation\<close>
inductive
- typing :: "cxt\<Rightarrow>llam\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [60,60,60] 60)
+ typing :: "cxt\<Rightarrow>llam\<Rightarrow>ty\<Rightarrow>bool" (\<open> _ \<turnstile> _ : _ \<close> [60,60,60] 60)
where
t_lPar[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> lPar x : T"
| t_lApp[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> lApp t1 t2 : T2"
@@ -126,7 +126,7 @@
text \<open>sub-contexts\<close>
abbreviation
- "sub_context" :: "cxt \<Rightarrow> cxt \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60)
+ "sub_context" :: "cxt \<Rightarrow> cxt \<Rightarrow> bool" (\<open>_ \<subseteq> _\<close> [60,60] 60)
where
"\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>1 \<longrightarrow> (x,T)\<in>set \<Gamma>2"
--- 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 \<rightharpoonup> 'b" ("(1[_])")
+ "_Map" :: "maplets => 'a \<rightharpoonup> 'b" (\<open>(1[_])\<close>)
atom_decl name
nominal_datatype ty =
Atom nat
- | Arrow ty ty (infixr "\<rightarrow>" 200)
- | TupleT ty ty (infixr "\<otimes>" 210)
+ | Arrow ty ty (infixr \<open>\<rightarrow>\<close> 200)
+ | TupleT ty ty (infixr \<open>\<otimes>\<close> 210)
lemma fresh_type [simp]: "(a::name) \<sharp> (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'\<langle>_,/ _'\<rangle>)")
+ | Tuple trm trm (\<open>(1'\<langle>_,/ _'\<rangle>)\<close>)
| Abs ty "\<guillemotleft>name\<guillemotright>trm"
- | App trm trm (infixl "\<cdot>" 200)
+ | App trm trm (infixl \<open>\<cdot>\<close> 200)
| Let ty trm btrm
and btrm =
Base trm
| Bind ty "\<guillemotleft>name\<guillemotright>btrm"
abbreviation
- Abs_syn :: "name \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" ("(3\<lambda>_:_./ _)" [0, 0, 10] 10)
+ Abs_syn :: "name \<Rightarrow> ty \<Rightarrow> trm \<Rightarrow> trm" (\<open>(3\<lambda>_:_./ _)\<close> [0, 0, 10] 10)
where
"\<lambda>x:T. t \<equiv> Abs T x t"
datatype pat =
PVar name ty
- | PTuple pat pat ("(1'\<langle>\<langle>_,/ _'\<rangle>\<rangle>)")
+ | PTuple pat pat (\<open>(1'\<langle>\<langle>_,/ _'\<rangle>\<rangle>)\<close>)
(* FIXME: The following should be done automatically by the nominal package *)
overloading pat_perm \<equiv> "perm :: name prm \<Rightarrow> pat \<Rightarrow> 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 \<Rightarrow> btrm \<Rightarrow> btrm" ("(3\<lambda>[_]./ _)" [0, 10] 10)
+primrec abs_pat :: "pat \<Rightarrow> btrm \<Rightarrow> btrm" (\<open>(3\<lambda>[_]./ _)\<close> [0, 10] 10)
where
"(\<lambda>[PVar x T]. t) = Bind T x t"
| "(\<lambda>[\<langle>\<langle>p, q\<rangle>\<rangle>]. t) = (\<lambda>[p]. \<lambda>[q]. t)"
@@ -143,7 +143,7 @@
type_synonym ctx = "(name \<times> ty) list"
inductive
- ptyping :: "pat \<Rightarrow> ty \<Rightarrow> ctx \<Rightarrow> bool" ("\<turnstile> _ : _ \<Rightarrow> _" [60, 60, 60] 60)
+ ptyping :: "pat \<Rightarrow> ty \<Rightarrow> ctx \<Rightarrow> bool" (\<open>\<turnstile> _ : _ \<Rightarrow> _\<close> [60, 60, 60] 60)
where
PVar: "\<turnstile> PVar x T : T \<Rightarrow> [(x, T)]"
| PTuple: "\<turnstile> p : T \<Rightarrow> \<Delta>\<^sub>1 \<Longrightarrow> \<turnstile> q : U \<Rightarrow> \<Delta>\<^sub>2 \<Longrightarrow> \<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> : T \<otimes> U \<Rightarrow> \<Delta>\<^sub>2 @ \<Delta>\<^sub>1"
@@ -168,16 +168,16 @@
by (induct \<Gamma>) (auto simp add: fresh_ctxt_set_eq [symmetric])
abbreviation
- "sub_ctx" :: "ctx \<Rightarrow> ctx \<Rightarrow> bool" ("_ \<sqsubseteq> _")
+ "sub_ctx" :: "ctx \<Rightarrow> ctx \<Rightarrow> bool" (\<open>_ \<sqsubseteq> _\<close>)
where
"\<Gamma>\<^sub>1 \<sqsubseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
abbreviation
- Let_syn :: "pat \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" ("(LET (_ =/ _)/ IN (_))" 10)
+ Let_syn :: "pat \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" (\<open>(LET (_ =/ _)/ IN (_))\<close> 10)
where
"LET p = t IN u \<equiv> Let (pat_type p) t (\<lambda>[p]. Base u)"
-inductive typing :: "ctx \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [60, 60, 60] 60)
+inductive typing :: "ctx \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [60, 60, 60] 60)
where
Var [intro]: "valid \<Gamma> \<Longrightarrow> (x, T) \<in> set \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
| Tuple [intro]: "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> u : U \<Longrightarrow> \<Gamma> \<turnstile> \<langle>t, u\<rangle> : T \<otimes> U"
@@ -268,7 +268,7 @@
qed auto
inductive
- match :: "pat \<Rightarrow> trm \<Rightarrow> (name \<times> trm) list \<Rightarrow> bool" ("\<turnstile> _ \<rhd> _ \<Rightarrow> _" [50, 50, 50] 50)
+ match :: "pat \<Rightarrow> trm \<Rightarrow> (name \<times> trm) list \<Rightarrow> bool" (\<open>\<turnstile> _ \<rhd> _ \<Rightarrow> _\<close> [50, 50, 50] 50)
where
PVar: "\<turnstile> PVar x T \<rhd> t \<Rightarrow> [(x, t)]"
| PProd: "\<turnstile> p \<rhd> t \<Rightarrow> \<theta> \<Longrightarrow> \<turnstile> q \<rhd> u \<Rightarrow> \<theta>' \<Longrightarrow> \<turnstile> \<langle>\<langle>p, q\<rangle>\<rangle> \<rhd> \<langle>t, u\<rangle> \<Rightarrow> \<theta> @ \<theta>'"
@@ -287,8 +287,8 @@
by (induct \<theta>) (auto simp add: eqvts)
nominal_primrec
- psubst :: "(name \<times> trm) list \<Rightarrow> trm \<Rightarrow> trm" ("_\<lparr>_\<rparr>" [95,0] 210)
- and psubstb :: "(name \<times> trm) list \<Rightarrow> btrm \<Rightarrow> btrm" ("_\<lparr>_\<rparr>\<^sub>b" [95,0] 210)
+ psubst :: "(name \<times> trm) list \<Rightarrow> trm \<Rightarrow> trm" (\<open>_\<lparr>_\<rparr>\<close> [95,0] 210)
+ and psubstb :: "(name \<times> trm) list \<Rightarrow> btrm \<Rightarrow> btrm" (\<open>_\<lparr>_\<rparr>\<^sub>b\<close> [95,0] 210)
where
"\<theta>\<lparr>Var x\<rparr> = (lookup \<theta> x)"
| "\<theta>\<lparr>t \<cdot> u\<rparr> = \<theta>\<lparr>t\<rparr> \<cdot> \<theta>\<lparr>u\<rparr>"
@@ -320,12 +320,12 @@
(simp_all add: eqvts fresh_bij)
abbreviation
- subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_\<mapsto>_]" [100,0,0] 100)
+ subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" (\<open>_[_\<mapsto>_]\<close> [100,0,0] 100)
where
"t[x\<mapsto>t'] \<equiv> [(x,t')]\<lparr>t\<rparr>"
abbreviation
- substb :: "btrm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> btrm" ("_[_\<mapsto>_]\<^sub>b" [100,0,0] 100)
+ substb :: "btrm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> btrm" (\<open>_[_\<mapsto>_]\<^sub>b\<close> [100,0,0] 100)
where
"t[x\<mapsto>t']\<^sub>b \<equiv> [(x,t')]\<lparr>t\<rparr>\<^sub>b"
@@ -525,7 +525,7 @@
lemmas match_type = match_type_aux [where \<Gamma>\<^sub>1="[]", simplified]
-inductive eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longmapsto> _" [60,60] 60)
+inductive eval :: "trm \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ \<longmapsto> _\<close> [60,60] 60)
where
TupleL: "t \<longmapsto> t' \<Longrightarrow> \<langle>t, u\<rangle> \<longmapsto> \<langle>t', u\<rangle>"
| TupleR: "u \<longmapsto> u' \<Longrightarrow> \<langle>t, u\<rangle> \<longmapsto> \<langle>t, u'\<rangle>"
--- 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\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80)
+ Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>\<beta> _\<close> [80,80] 80)
where
b1[intro!]: "s1 \<longrightarrow>\<^sub>\<beta> s2 \<Longrightarrow> App s1 t \<longrightarrow>\<^sub>\<beta> App s2 t"
| b2[intro!]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^sub>\<beta> App t s2"
@@ -113,7 +113,7 @@
nominal_datatype ty =
TVar "nat"
- | TArr "ty" "ty" (infix "\<rightarrow>" 200)
+ | TArr "ty" "ty" (infix \<open>\<rightarrow>\<close> 200)
lemma fresh_ty:
fixes a ::"name"
@@ -144,7 +144,7 @@
(auto simp add: fresh_prod fresh_list_cons fresh_atm)
inductive
- typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
+ typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" (\<open>_ \<turnstile> _ : _\<close> [60,60,60] 60)
where
t1[intro]: "\<lbrakk>valid \<Gamma>; (a,\<tau>)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var a : \<tau>"
| t2[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>"
@@ -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\<Rightarrow>lam\<Rightarrow>bool" (" _ \<guillemotright> _" [80,80] 80)
+ FST :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<guillemotright> _\<close> [80,80] 80)
where
fst[intro!]: "(App t s) \<guillemotright> t"
@@ -455,12 +455,12 @@
qed
abbreviation
- mapsto :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55)
+ mapsto :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> bool" (\<open>_ maps _ to _\<close> [55,55,55] 55)
where
"\<theta> maps x to e \<equiv> (lookup \<theta> x) = e"
abbreviation
- closes :: "(name\<times>lam) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ closes _" [55,55] 55)
+ closes :: "(name\<times>lam) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (\<open>_ closes _\<close> [55,55] 55)
where
"\<theta> closes \<Gamma> \<equiv> \<forall>x T. ((x,T) \<in> set \<Gamma> \<longrightarrow> (\<exists>t. \<theta> maps x to t \<and> t \<in> RED T))"
--- 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 \<open>types and terms\<close>
nominal_datatype ty =
TVar "nat"
- | Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
+ | Arrow "ty" "ty" (\<open>_\<rightarrow>_\<close> [100,100] 100)
nominal_datatype trm =
Var "name"
- | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>trm" (\<open>Lam [_]._\<close> [100,100] 100)
| App "trm" "trm"
lemma fresh_ty:
@@ -62,7 +62,7 @@
(* parallel substitution *)
nominal_primrec
- psubst :: "(name\<times>trm) list \<Rightarrow> trm \<Rightarrow> trm" ("_<_>" [95,95] 105)
+ psubst :: "(name\<times>trm) list \<Rightarrow> trm \<Rightarrow> trm" (\<open>_<_>\<close> [95,95] 105)
where
"\<theta><(Var x)> = (lookup \<theta> x)"
| "\<theta><(App e\<^sub>1 e\<^sub>2)> = App (\<theta><e\<^sub>1>) (\<theta><e\<^sub>2>)"
@@ -92,7 +92,7 @@
text \<open>Single substitution\<close>
abbreviation
- subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
+ subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" (\<open>_[_::=_]\<close> [100,100,100] 100)
where
"t[x::=t'] \<equiv> ([(x,t')])<t>"
@@ -157,7 +157,7 @@
text \<open>Typing Relation\<close>
inductive
- typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
+ typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (\<open>_ \<turnstile> _ : _\<close> [60,60,60] 60)
where
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> e\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> e\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> 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\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [55,55] 55)
+ "sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (\<open>_ \<subseteq> _\<close> [55,55] 55)
where
"\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>\<^sub>1 \<longrightarrow> (x,T)\<in>set \<Gamma>\<^sub>2"
@@ -257,7 +257,7 @@
text \<open>Big-Step Evaluation\<close>
inductive
- big :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80)
+ big :: "trm\<Rightarrow>trm\<Rightarrow>bool" (\<open>_ \<Down> _\<close> [80,80] 80)
where
b_Lam[intro]: "Lam [x].e \<Down> Lam [x].e"
| b_App[intro]: "\<lbrakk>x\<sharp>(e\<^sub>1,e\<^sub>2,e'); e\<^sub>1\<Down>Lam [x].e; e\<^sub>2\<Down>e\<^sub>2'; e[x::=e\<^sub>2']\<Down>e'\<rbrakk> \<Longrightarrow> App e\<^sub>1 e\<^sub>2 \<Down> e'"
@@ -427,12 +427,12 @@
text \<open>'\<theta> maps x to e' asserts that \<theta> substitutes x with e\<close>
abbreviation
- mapsto :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55)
+ mapsto :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> bool" (\<open>_ maps _ to _\<close> [55,55,55] 55)
where
"\<theta> maps x to e \<equiv> (lookup \<theta> x) = e"
abbreviation
- v_closes :: "(name\<times>trm) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ Vcloses _" [55,55] 55)
+ v_closes :: "(name\<times>trm) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (\<open>_ Vcloses _\<close> [55,55] 55)
where
"\<theta> Vcloses \<Gamma> \<equiv> \<forall>x T. (x,T) \<in> set \<Gamma> \<longrightarrow> (\<exists>v. \<theta> maps x to v \<and> v \<in> V T)"
--- 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 "\<degree>" 200)
-| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [0, 10] 10)
+| App "lam" "lam" (infixl \<open>\<degree>\<close> 200)
+| Lam "\<guillemotleft>name\<guillemotright>lam" (\<open>Lam [_]._\<close> [0, 10] 10)
instantiation lam :: size
begin
@@ -38,7 +38,7 @@
end
nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [300, 0, 0] 300)
+ subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" (\<open>_[_::=_]\<close> [300, 0, 0] 300)
where
subst_Var: "(Var x)[y::=s] = (if x=y then s else (Var x))"
| subst_App: "(t\<^sub>1 \<degree> t\<^sub>2)[y::=s] = t\<^sub>1[y::=s] \<degree> t\<^sub>2[y::=s]"
@@ -82,7 +82,7 @@
lemma subst_neq [simp]: "x \<noteq> y \<Longrightarrow> (Var x)[y::=u] = Var x"
by (simp add: subst_Var)
-inductive beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+inductive beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<close> 50)
where
beta: "x \<sharp> t \<Longrightarrow> (Lam [x].s) \<degree> t \<rightarrow>\<^sub>\<beta> s[x::=t]"
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
@@ -106,14 +106,14 @@
qed
abbreviation
- beta_reds :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>\<beta>\<^sup>*" 50) where
+ beta_reds :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<^sup>*\<close> 50) where
"s \<rightarrow>\<^sub>\<beta>\<^sup>* t \<equiv> beta\<^sup>*\<^sup>* s t"
subsection \<open>Application of a term to a list of terms\<close>
abbreviation
- list_application :: "lam \<Rightarrow> lam list \<Rightarrow> lam" (infixl "\<degree>\<degree>" 150) where
+ list_application :: "lam \<Rightarrow> lam list \<Rightarrow> lam" (infixl \<open>\<degree>\<degree>\<close> 150) where
"t \<degree>\<degree> ts \<equiv> foldl (\<degree>) t ts"
lemma apps_eq_tail_conv [iff]: "(r \<degree>\<degree> ts = s \<degree>\<degree> ts) = (r = s)"
@@ -318,7 +318,7 @@
subsection \<open>Lifting beta-reduction to lists\<close>
abbreviation
- list_beta :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>\<beta>]\<^sub>1" 50) where
+ list_beta :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl \<open>[\<rightarrow>\<^sub>\<beta>]\<^sub>1\<close> 50) where
"rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<equiv> 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 \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>s" 50)
- and sredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>s]" 50)
+ sred :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl \<open>\<rightarrow>\<^sub>s\<close> 50)
+ and sredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl \<open>[\<rightarrow>\<^sub>s]\<close> 50)
where
"s [\<rightarrow>\<^sub>s] t \<equiv> listrelp (\<rightarrow>\<^sub>s) s t"
| Var: "rs [\<rightarrow>\<^sub>s] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> rs'"
@@ -697,8 +697,8 @@
subsection \<open>Leftmost reduction and weakly normalizing terms\<close>
inductive
- lred :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>l" 50)
- and lredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>l]" 50)
+ lred :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl \<open>\<rightarrow>\<^sub>l\<close> 50)
+ and lredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl \<open>[\<rightarrow>\<^sub>l]\<close> 50)
where
"s [\<rightarrow>\<^sub>l] t \<equiv> listrelp (\<rightarrow>\<^sub>l) s t"
| Var: "rs [\<rightarrow>\<^sub>l] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>l Var x \<degree>\<degree> rs'"
--- 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 "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._")
+| Lam "\<guillemotleft>name\<guillemotright>lam" (\<open>Lam [_]._\<close>)
text \<open>Capture-Avoiding Substitution\<close>
nominal_primrec
- subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]")
+ subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" (\<open>_[_::=_]\<close>)
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" ("_ \<rightarrow> _")
+ | TArr "ty" "ty" (\<open>_ \<rightarrow> _\<close>)
lemma ty_fresh:
fixes x::"name"
@@ -61,7 +61,7 @@
type_synonym ctx = "(name\<times>ty) list"
abbreviation
- "sub_ctx" :: "ctx \<Rightarrow> ctx \<Rightarrow> bool" ("_ \<subseteq> _")
+ "sub_ctx" :: "ctx \<Rightarrow> ctx \<Rightarrow> bool" (\<open>_ \<subseteq> _\<close>)
where
"\<Gamma>\<^sub>1 \<subseteq> \<Gamma>\<^sub>2 \<equiv> \<forall>x. x \<in> set \<Gamma>\<^sub>1 \<longrightarrow> x \<in> set \<Gamma>\<^sub>2"
@@ -102,7 +102,7 @@
text \<open>Typing Relation\<close>
inductive
- typing :: "ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _")
+ typing :: "ctx \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close>)
where
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t\<^sub>1 t\<^sub>2 : T\<^sub>2"
@@ -169,7 +169,7 @@
text \<open>Beta Reduction\<close>
inductive
- "beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _")
+ "beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>\<beta> _\<close>)
where
b1[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> App t1 s \<longrightarrow>\<^sub>\<beta> App t2 s"
| b2[intro]: "s1 \<longrightarrow>\<^sub>\<beta> s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^sub>\<beta> App t s2"
--- 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 "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>lam" (\<open>Lam [_]._\<close> [100,100] 100)
text \<open>
The inductive relation 'unbind' unbinds the top-most
@@ -27,7 +27,7 @@
of the algorithm W.\<close>
inductive
- unbind :: "lam \<Rightarrow> name list \<Rightarrow> lam \<Rightarrow> bool" ("_ \<mapsto> _,_" [60,60,60] 60)
+ unbind :: "lam \<Rightarrow> name list \<Rightarrow> lam \<Rightarrow> bool" (\<open>_ \<mapsto> _,_\<close> [60,60,60] 60)
where
u_var: "(Var a) \<mapsto> [],(Var a)"
| u_app: "(App t1 t2) \<mapsto> [],(App t1 t2)"
@@ -143,7 +143,7 @@
strips off the top-most binders from lambdas.\<close>
inductive
- strip :: "lam \<Rightarrow> lam \<Rightarrow> bool" ("_ \<rightarrow> _" [60,60] 60)
+ strip :: "lam \<Rightarrow> lam \<Rightarrow> bool" (\<open>_ \<rightarrow> _\<close> [60,60] 60)
where
s_var: "(Var a) \<rightarrow> (Var a)"
| s_app: "(App t1 t2) \<rightarrow> (App t1 t2)"
--- 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 \<Rightarrow> 'a list \<Rightarrow> 'a list" ("_ - _" [60,60] 60)
+ "difference_list" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (\<open>_ - _\<close> [60,60] 60)
where
"xs - ys \<equiv> [x \<leftarrow> xs. x\<notin>set ys]"
@@ -34,20 +34,20 @@
nominal_datatype ty =
TVar "tvar"
- | Fun "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
+ | Fun "ty" "ty" (\<open>_\<rightarrow>_\<close> [100,100] 100)
nominal_datatype tyS =
Ty "ty"
- | ALL "\<guillemotleft>tvar\<guillemotright>tyS" ("\<forall>[_]._" [100,100] 100)
+ | ALL "\<guillemotleft>tvar\<guillemotright>tyS" (\<open>\<forall>[_]._\<close> [100,100] 100)
nominal_datatype trm =
Var "var"
| App "trm" "trm"
- | Lam "\<guillemotleft>var\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>var\<guillemotright>trm" (\<open>Lam [_]._\<close> [100,100] 100)
| Let "\<guillemotleft>var\<guillemotright>trm" "trm"
abbreviation
- LetBe :: "var \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" ("Let _ be _ in _" [100,100,100] 100)
+ LetBe :: "var \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" (\<open>Let _ be _ in _\<close> [100,100,100] 100)
where
"Let x be t1 in t2 \<equiv> trm.Let x t2 t1"
@@ -175,10 +175,10 @@
type_synonym Subst = "(tvar\<times>ty) list"
consts
- psubst :: "Subst \<Rightarrow> 'a \<Rightarrow> 'a" ("_<_>" [100,60] 120)
+ psubst :: "Subst \<Rightarrow> 'a \<Rightarrow> 'a" (\<open>_<_>\<close> [100,60] 120)
abbreviation
- subst :: "'a \<Rightarrow> tvar \<Rightarrow> ty \<Rightarrow> 'a" ("_[_::=_]" [100,100,100] 100)
+ subst :: "'a \<Rightarrow> tvar \<Rightarrow> ty \<Rightarrow> 'a" (\<open>_[_::=_]\<close> [100,100,100] 100)
where
"smth[X::=T] \<equiv> ([(X,T)])<smth>"
@@ -299,7 +299,7 @@
text \<open>instance of a type scheme\<close>
inductive
- inst :: "ty \<Rightarrow> tyS \<Rightarrow> bool"("_ \<prec> _" [50,51] 50)
+ inst :: "ty \<Rightarrow> tyS \<Rightarrow> bool"(\<open>_ \<prec> _\<close> [50,51] 50)
where
I_Ty[intro]: "T \<prec> (Ty T)"
| I_All[intro]: "\<lbrakk>X\<sharp>T'; T \<prec> S\<rbrakk> \<Longrightarrow> T[X::=T'] \<prec> \<forall>[X].S"
@@ -350,7 +350,7 @@
text\<open>typing judgements\<close>
inductive
- typing :: "Ctxt \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (" _ \<turnstile> _ : _ " [60,60,60] 60)
+ typing :: "Ctxt \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (\<open> _ \<turnstile> _ : _ \<close> [60,60,60] 60)
where
T_VAR[intro]: "\<lbrakk>valid \<Gamma>; (x,S)\<in>set \<Gamma>; T \<prec> S\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
| T_APP[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>1\<rightarrow>T\<^sub>2; \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t\<^sub>1 t\<^sub>2 : T\<^sub>2"
--- 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 "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
+ | Lam "\<guillemotleft>name\<guillemotright>lam" (\<open>Lam [_]._\<close> [100,100] 100)
nominal_datatype ty =
TVar "string"
- | TArr "ty" "ty" ("_ \<rightarrow> _" [100,100] 100)
+ | TArr "ty" "ty" (\<open>_ \<rightarrow> _\<close> [100,100] 100)
lemma ty_fresh:
fixes x::"name"
@@ -42,7 +42,7 @@
text\<open>Typing judgements\<close>
inductive
- typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
+ typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" (\<open>_ \<turnstile> _ : _\<close> [60,60,60] 60)
where
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
@@ -60,7 +60,7 @@
text \<open>Abbreviation for the notion of subcontexts.\<close>
abbreviation
- "sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60)
+ "sub_context" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (\<open>_ \<subseteq> _\<close> [60,60] 60)
where
"\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>x T. (x,T)\<in>set \<Gamma>1 \<longrightarrow> (x,T)\<in>set \<Gamma>2"
--- 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 \<open>Nonstandard Definitions.\<close>
definition nsderiv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool"
- ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
+ (\<open>(NSDERIV (_)/ (_)/ :> (_))\<close> [1000, 1000, 60] 60)
where "NSDERIV f x :> D \<longleftrightarrow>
(\<forall>h \<in> Infinitesimal - {0}. (( *f* f)(star_of x + h) - star_of (f x)) / h \<approx> star_of D)"
definition NSdifferentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
- (infixl "NSdifferentiable" 60)
+ (infixl \<open>NSdifferentiable\<close> 60)
where "f NSdifferentiable x \<longleftrightarrow> (\<exists>D. NSDERIV f x :> D)"
definition increment :: "(real \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> hypreal \<Rightarrow> hypreal"
--- 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 \<open>Nonstandard Definitions.\<close>
definition NSLIM :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
- ("((_)/ \<midarrow>(_)/\<rightarrow>\<^sub>N\<^sub>S (_))" [60, 0, 60] 60)
+ (\<open>((_)/ \<midarrow>(_)/\<rightarrow>\<^sub>N\<^sub>S (_))\<close> [60, 0, 60] 60)
where "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<forall>x. x \<noteq> star_of a \<and> x \<approx> star_of a \<longrightarrow> ( *f* f) x \<approx> star_of L)"
definition isNSCont :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
--- 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 \<Rightarrow> hypreal \<Rightarrow> hypreal" (infixr "powhr" 80)
+definition powhr :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hypreal" (infixr \<open>powhr\<close> 80)
where [transfer_unfold]: "x powhr a = starfun2 (powr) x a"
definition hlog :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hypreal"
--- 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 \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
- ("((_)/ \<longlonglongrightarrow>\<^sub>N\<^sub>S (_))" [60, 60] 60) where
+ (\<open>((_)/ \<longlonglongrightarrow>\<^sub>N\<^sub>S (_))\<close> [60, 60] 60) where
\<comment> \<open>Nonstandard definition of convergence of sequence\<close>
"X \<longlonglongrightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
--- 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 \<times> hypnat \<times> (nat \<Rightarrow> real) \<Rightarrow> hypreal"
where "sumhr = (\<lambda>(M,N,f). starfun2 (\<lambda>m n. sum f {m..<n}) M N)"
-definition NSsums :: "(nat \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> bool" (infixr "NSsums" 80)
+definition NSsums :: "(nat \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> bool" (infixr \<open>NSsums\<close> 80)
where "f NSsums s = (\<lambda>n. sum f {..<n}) \<longlonglongrightarrow>\<^sub>N\<^sub>S s"
definition NSsummable :: "(nat \<Rightarrow> real) \<Rightarrow> bool"
--- 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 \<Rightarrow> hypreal"
where "hypreal_of_hypnat \<equiv> of_hypnat"
-definition omega :: hypreal ("\<omega>")
+definition omega :: hypreal (\<open>\<omega>\<close>)
where "\<omega> = star_n (\<lambda>n. real (Suc n))"
\<comment> \<open>an infinite number \<open>= [<1, 2, 3, \<dots>>]\<close>\<close>
-definition epsilon :: hypreal ("\<epsilon>")
+definition epsilon :: hypreal (\<open>\<epsilon>\<close>)
where "\<epsilon> = star_n (\<lambda>n. inverse (real (Suc n)))"
\<comment> \<open>an infinitesimal number \<open>= [<1, 1/2, 1/3, \<dots>>]\<close>\<close>
@@ -301,7 +301,7 @@
subsection \<open>Powers with Hypernatural Exponents\<close>
text \<open>Hypernatural powers of hyperreals.\<close>
-definition pow :: "'a::power star \<Rightarrow> nat star \<Rightarrow> 'a star" (infixr "pow" 80)
+definition pow :: "'a::power star \<Rightarrow> nat star \<Rightarrow> 'a star" (infixr \<open>pow\<close> 80)
where hyperpow_def [transfer_unfold]: "R pow N = ( *f2* (^)) R N"
lemma Standard_hyperpow [simp]: "r \<in> Standard \<Longrightarrow> n \<in> Standard \<Longrightarrow> r pow n \<in> Standard"
--- 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. \<forall>r \<in> Reals. r < hnorm x}"
-definition approx :: "'a::real_normed_vector star \<Rightarrow> 'a star \<Rightarrow> bool" (infixl "\<approx>" 50)
+definition approx :: "'a::real_normed_vector star \<Rightarrow> 'a star \<Rightarrow> bool" (infixl \<open>\<approx>\<close> 50)
where "x \<approx> y \<longleftrightarrow> x - y \<in> Infinitesimal"
\<comment> \<open>the ``infinitely close'' relation\<close>
--- 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 \<comment> \<open>internal sets\<close>
- starset_n :: "(nat \<Rightarrow> 'a set) \<Rightarrow> 'a star set" ("*sn* _" [80] 80)
+ starset_n :: "(nat \<Rightarrow> 'a set) \<Rightarrow> 'a star set" (\<open>*sn* _\<close> [80] 80)
where "*sn* As = Iset (star_n As)"
definition InternalSets :: "'a star set set"
@@ -23,7 +23,7 @@
(\<forall>x y. \<exists>X \<in> Rep_star x. \<exists>Y \<in> Rep_star y. y = F x \<longleftrightarrow> eventually (\<lambda>n. Y n = f(X n)) \<U>)"
definition \<comment> \<open>internal functions\<close>
- starfun_n :: "(nat \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a star \<Rightarrow> 'b star" ("*fn* _" [80] 80)
+ starfun_n :: "(nat \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a star \<Rightarrow> 'b star" (\<open>*fn* _\<close> [80] 80)
where "*fn* F = Ifun (star_n F)"
definition InternalFuns :: "('a star => 'b star) set"
--- 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 \<Rightarrow> 'a \<Rightarrow> bool" (infixl "dvd'_strict" 50)
+abbreviation dvd_strict :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>dvd'_strict\<close> 50)
where
"b dvd_strict a \<equiv> b dvd a \<and> \<not> a dvd b"
--- 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 "\<Rightarrow>" 200)
+ | Fun type type (infixr \<open>\<Rightarrow>\<close> 200)
datatype dB =
Var nat
- | App dB dB (infixl "\<degree>" 200)
+ | App dB dB (infixl \<open>\<degree>\<close> 200)
| Abs type dB
primrec
- nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+ nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" (\<open>_\<langle>_\<rangle>\<close> [90, 0] 91)
where
"[]\<langle>i\<rangle> = None"
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
@@ -184,7 +184,7 @@
"nth_el' (x # xs) 0 x"
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
-inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [50, 50, 50] 50)
where
Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
| Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> 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" (\<open>_[_'/_]\<close> [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 \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
| subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
-inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+inductive beta :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<close> 50)
where
beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> 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 (\<open>_ ::= _\<close> [1000, 61] 61)
+ | Semi com' com' (\<open>_; _\<close> [60, 61] 60)
+ | If bexp com' com' (\<open>IF _ THEN _ ELSE _\<close> [0, 0, 61] 61)
+ | While bexp com' (\<open>WHILE _ DO _\<close> [0, 61] 61)
inductive
- big_step :: "com' * state' \<Rightarrow> state' \<Rightarrow> bool" (infix "\<Rightarrow>" 55)
+ big_step :: "com' * state' \<Rightarrow> state' \<Rightarrow> bool" (infix \<open>\<Rightarrow>\<close> 55)
where
Skip: "(SKIP,s) \<Rightarrow> s"
| Assign: "(x ::= a,s) \<Rightarrow> s(x := aval a s)"
--- 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 "\<Rightarrow>" 200)
+ | Fun type type (infixr \<open>\<Rightarrow>\<close> 200)
datatype dB =
Var nat
- | App dB dB (infixl "\<degree>" 200)
+ | App dB dB (infixl \<open>\<degree>\<close> 200)
| Abs type dB
primrec
- nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+ nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" (\<open>_\<langle>_\<rangle>\<close> [90, 0] 91)
where
"[]\<langle>i\<rangle> = None"
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
@@ -24,7 +24,7 @@
"nth_el1 (x # xs) 0 x"
| "nth_el1 xs i y \<Longrightarrow> nth_el1 (x # xs) (Suc i) y"
-inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [50, 50, 50] 50)
where
Var [intro!]: "nth_el1 env x T \<Longrightarrow> env \<turnstile> Var x : T"
| Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
@@ -42,14 +42,14 @@
"pred (Suc i) = i"
primrec
- subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300)
+ subst :: "[dB, dB, nat] => dB" (\<open>_[_'/_]\<close> [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 \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
| subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
-inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+inductive beta :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<close> 50)
where
beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
--- 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 "\<Rightarrow>" 200)
+ | Fun type type (infixr \<open>\<Rightarrow>\<close> 200)
datatype dB =
Var nat
- | App dB dB (infixl "\<degree>" 200)
+ | App dB dB (infixl \<open>\<degree>\<close> 200)
| Abs type dB
primrec
- nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+ nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" (\<open>_\<langle>_\<rangle>\<close> [90, 0] 91)
where
"[]\<langle>i\<rangle> = None"
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
@@ -228,7 +228,7 @@
"nth_el' (x # xs) 0 x"
| "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
-inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [50, 50, 50] 50)
where
Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
| Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> 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" (\<open>_[_'/_]\<close> [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 \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
| subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
-inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+inductive beta :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<close> 50)
where
beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
--- 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 \<open>Specialisation in POPLmark theory\<close>
notation
- Some ("\<lfloor>_\<rfloor>")
+ Some (\<open>\<lfloor>_\<rfloor>\<close>)
notation
- None ("\<bottom>")
+ None (\<open>\<bottom>\<close>)
notation
- length ("\<parallel>_\<parallel>")
+ length (\<open>\<parallel>_\<parallel>\<close>)
notation
- Cons ("_ \<Colon>/ _" [66, 65] 65)
+ Cons (\<open>_ \<Colon>/ _\<close> [66, 65] 65)
primrec
- nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
+ nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" (\<open>_\<langle>_\<rangle>\<close> [90, 0] 91)
where
"[]\<langle>i\<rangle> = \<bottom>"
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> \<lfloor>x\<rfloor> | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
-primrec assoc :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b option" ("_\<langle>_\<rangle>\<^sub>?" [90, 0] 91)
+primrec assoc :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b option" (\<open>_\<langle>_\<rangle>\<^sub>?\<close> [90, 0] 91)
where
"[]\<langle>a\<rangle>\<^sub>? = \<bottom>"
| "(x # xs)\<langle>a\<rangle>\<^sub>? = (if fst x = a then \<lfloor>snd x\<rfloor> else xs\<langle>a\<rangle>\<^sub>?)"
@@ -125,8 +125,8 @@
datatype type =
TVar nat
| Top
- | Fun type type (infixr "\<rightarrow>" 200)
- | TyAll type type ("(3\<forall><:_./ _)" [0, 10] 10)
+ | Fun type type (infixr \<open>\<rightarrow>\<close> 200)
+ | TyAll type type (\<open>(3\<forall><:_./ _)\<close> [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\<lambda>:_./ _)" [0, 10] 10)
- | TAbs type trm ("(3\<lambda><:_./ _)" [0, 10] 10)
- | App trm trm (infixl "\<bullet>" 200)
- | TApp trm type (infixl "\<bullet>\<^sub>\<tau>" 200)
+ | Abs type trm (\<open>(3\<lambda>:_./ _)\<close> [0, 10] 10)
+ | TAbs type trm (\<open>(3\<lambda><:_./ _)\<close> [0, 10] 10)
+ | App trm trm (infixl \<open>\<bullet>\<close> 200)
+ | TApp trm type (infixl \<open>\<bullet>\<^sub>\<tau>\<close> 200)
-primrec liftT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("\<up>\<^sub>\<tau>")
+primrec liftT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" (\<open>\<up>\<^sub>\<tau>\<close>)
where
"\<up>\<^sub>\<tau> n k (TVar i) = (if i < k then TVar i else TVar (i + n))"
| "\<up>\<^sub>\<tau> n k Top = Top"
| "\<up>\<^sub>\<tau> n k (T \<rightarrow> U) = \<up>\<^sub>\<tau> n k T \<rightarrow> \<up>\<^sub>\<tau> n k U"
| "\<up>\<^sub>\<tau> n k (\<forall><:T. U) = (\<forall><:\<up>\<^sub>\<tau> n k T. \<up>\<^sub>\<tau> n (k + 1) U)"
-primrec lift :: "nat \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm" ("\<up>")
+primrec lift :: "nat \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm" (\<open>\<up>\<close>)
where
"\<up> n k (Var i) = (if i < k then Var i else Var (i + n))"
| "\<up> n k (\<lambda>:T. t) = (\<lambda>:\<up>\<^sub>\<tau> n k T. \<up> n (k + 1) t)"
@@ -168,7 +168,7 @@
| "\<up> n k (s \<bullet> t) = \<up> n k s \<bullet> \<up> n k t"
| "\<up> n k (t \<bullet>\<^sub>\<tau> T) = \<up> n k t \<bullet>\<^sub>\<tau> \<up>\<^sub>\<tau> n k T"
-primrec substTT :: "type \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("_[_ \<mapsto>\<^sub>\<tau> _]\<^sub>\<tau>" [300, 0, 0] 300)
+primrec substTT :: "type \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" (\<open>_[_ \<mapsto>\<^sub>\<tau> _]\<^sub>\<tau>\<close> [300, 0, 0] 300)
where
"(TVar i)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> =
(if k < i then TVar (i - 1) else if i = k then \<up>\<^sub>\<tau> k 0 S else TVar i)"
@@ -176,12 +176,12 @@
| "(T \<rightarrow> U)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> = T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> \<rightarrow> U[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>"
| "(\<forall><:T. U)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> = (\<forall><:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. U[k+1 \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>)"
-primrec decT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("\<down>\<^sub>\<tau>")
+primrec decT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" (\<open>\<down>\<^sub>\<tau>\<close>)
where
"\<down>\<^sub>\<tau> 0 k T = T"
| "\<down>\<^sub>\<tau> (Suc n) k T = \<down>\<^sub>\<tau> n k (T[k \<mapsto>\<^sub>\<tau> Top]\<^sub>\<tau>)"
-primrec subst :: "trm \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm" ("_[_ \<mapsto> _]" [300, 0, 0] 300)
+primrec subst :: "trm \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm" (\<open>_[_ \<mapsto> _]\<close> [300, 0, 0] 300)
where
"(Var i)[k \<mapsto> s] = (if k < i then Var (i - 1) else if i = k then \<up> k 0 s else Var i)"
| "(t \<bullet> u)[k \<mapsto> s] = t[k \<mapsto> s] \<bullet> u[k \<mapsto> s]"
@@ -189,7 +189,7 @@
| "(\<lambda>:T. t)[k \<mapsto> s] = (\<lambda>:\<down>\<^sub>\<tau> 1 k T. t[k+1 \<mapsto> s])"
| "(\<lambda><:T. t)[k \<mapsto> s] = (\<lambda><:\<down>\<^sub>\<tau> 1 k T. t[k+1 \<mapsto> s])"
-primrec substT :: "trm \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> trm" ("_[_ \<mapsto>\<^sub>\<tau> _]" [300, 0, 0] 300)
+primrec substT :: "trm \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> trm" (\<open>_[_ \<mapsto>\<^sub>\<tau> _]\<close> [300, 0, 0] 300)
where
"(Var i)[k \<mapsto>\<^sub>\<tau> S] = (if k < i then Var (i - 1) else Var i)"
| "(t \<bullet> u)[k \<mapsto>\<^sub>\<tau> S] = t[k \<mapsto>\<^sub>\<tau> S] \<bullet> u[k \<mapsto>\<^sub>\<tau> S]"
@@ -197,23 +197,23 @@
| "(\<lambda>:T. t)[k \<mapsto>\<^sub>\<tau> S] = (\<lambda>:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. t[k+1 \<mapsto>\<^sub>\<tau> S])"
| "(\<lambda><:T. t)[k \<mapsto>\<^sub>\<tau> S] = (\<lambda><:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. t[k+1 \<mapsto>\<^sub>\<tau> S])"
-primrec liftE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" ("\<up>\<^sub>e")
+primrec liftE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" (\<open>\<up>\<^sub>e\<close>)
where
"\<up>\<^sub>e n k [] = []"
| "\<up>\<^sub>e n k (B \<Colon> \<Gamma>) = mapB (\<up>\<^sub>\<tau> n (k + \<parallel>\<Gamma>\<parallel>)) B \<Colon> \<up>\<^sub>e n k \<Gamma>"
-primrec substE :: "env \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> env" ("_[_ \<mapsto>\<^sub>\<tau> _]\<^sub>e" [300, 0, 0] 300)
+primrec substE :: "env \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> env" (\<open>_[_ \<mapsto>\<^sub>\<tau> _]\<^sub>e\<close> [300, 0, 0] 300)
where
"[][k \<mapsto>\<^sub>\<tau> T]\<^sub>e = []"
| "(B \<Colon> \<Gamma>)[k \<mapsto>\<^sub>\<tau> T]\<^sub>e = mapB (\<lambda>U. U[k + \<parallel>\<Gamma>\<parallel> \<mapsto>\<^sub>\<tau> T]\<^sub>\<tau>) B \<Colon> \<Gamma>[k \<mapsto>\<^sub>\<tau> T]\<^sub>e"
-primrec decE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" ("\<down>\<^sub>e")
+primrec decE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" (\<open>\<down>\<^sub>e\<close>)
where
"\<down>\<^sub>e 0 k \<Gamma> = \<Gamma>"
| "\<down>\<^sub>e (Suc n) k \<Gamma> = \<down>\<^sub>e n k (\<Gamma>[k \<mapsto>\<^sub>\<tau> Top]\<^sub>e)"
inductive
- well_formed :: "env \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile>\<^sub>w\<^sub>f _" [50, 50] 50)
+ well_formed :: "env \<Rightarrow> type \<Rightarrow> bool" (\<open>_ \<turnstile>\<^sub>w\<^sub>f _\<close> [50, 50] 50)
where
wf_TVar: "\<Gamma>\<langle>i\<rangle> = \<lfloor>TVarB T\<rfloor> \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f TVar i"
| wf_Top: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f Top"
@@ -221,8 +221,8 @@
| wf_all: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f T \<Longrightarrow> TVarB T \<Colon> \<Gamma> \<turnstile>\<^sub>w\<^sub>f U \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (\<forall><:T. U)"
inductive
- well_formedE :: "env \<Rightarrow> bool" ("_ \<turnstile>\<^sub>w\<^sub>f" [50] 50)
- and well_formedB :: "env \<Rightarrow> binding \<Rightarrow> bool" ("_ \<turnstile>\<^sub>w\<^sub>f\<^sub>B _" [50, 50] 50)
+ well_formedE :: "env \<Rightarrow> bool" (\<open>_ \<turnstile>\<^sub>w\<^sub>f\<close> [50] 50)
+ and well_formedB :: "env \<Rightarrow> binding \<Rightarrow> bool" (\<open>_ \<turnstile>\<^sub>w\<^sub>f\<^sub>B _\<close> [50, 50] 50)
where
"\<Gamma> \<turnstile>\<^sub>w\<^sub>f\<^sub>B B \<equiv> \<Gamma> \<turnstile>\<^sub>w\<^sub>f type_ofB B"
| wf_Nil: "[] \<turnstile>\<^sub>w\<^sub>f"
@@ -238,7 +238,7 @@
"B \<Colon> \<Gamma> \<turnstile>\<^sub>w\<^sub>f"
inductive
- subtyping :: "env \<Rightarrow> type \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ <: _" [50, 50, 50] 50)
+ subtyping :: "env \<Rightarrow> type \<Rightarrow> type \<Rightarrow> bool" (\<open>_ \<turnstile> _ <: _\<close> [50, 50, 50] 50)
where
SA_Top: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f S \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"
| SA_refl_TVar: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f TVar i \<Longrightarrow> \<Gamma> \<turnstile> TVar i <: TVar i"
@@ -249,7 +249,7 @@
\<Gamma> \<turnstile> (\<forall><:S\<^sub>1. S\<^sub>2) <: (\<forall><:T\<^sub>1. T\<^sub>2)"
inductive
- typing :: "env \<Rightarrow> trm \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+ typing :: "env \<Rightarrow> trm \<Rightarrow> type \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [50, 50, 50] 50)
where
T_Var: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Longrightarrow> \<Gamma>\<langle>i\<rangle> = \<lfloor>VarB U\<rfloor> \<Longrightarrow> T = \<up>\<^sub>\<tau> (Suc i) 0 U \<Longrightarrow> \<Gamma> \<turnstile> Var i : T"
| T_Abs: "VarB T\<^sub>1 \<Colon> \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>:T\<^sub>1. t\<^sub>2) : T\<^sub>1 \<rightarrow> \<down>\<^sub>\<tau> 1 0 T\<^sub>2"
--- 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 \<Rightarrow> 'a measure \<Rightarrow> 'a measure" (infix "\<star>" 50) where
+definition convolution :: "('a :: ordered_euclidean_space) measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" (infix \<open>\<star>\<close> 50) where
"convolution M N = distr (M \<Otimes>\<^sub>M N) borel (\<lambda>(x, y). x + y)"
lemma
--- 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>\<open>Pi\<^sub>E\<close> and their sigma-algebra
\<^const>\<open>Pi\<^sub>M\<close>.\<close>
-type_notation fmap ("(_ \<Rightarrow>\<^sub>F /_)" [22, 21] 21)
+type_notation fmap (\<open>(_ \<Rightarrow>\<^sub>F /_)\<close> [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 \<Rightarrow>\<^sub>F 'a) \<Rightarrow> 'i \<Rightarrow> 'a" ("'((_)')\<^sub>F" [0] 1000) is
+lift_definition proj :: "('i \<Rightarrow>\<^sub>F 'a) \<Rightarrow> 'i \<Rightarrow> 'a" (\<open>'((_)')\<^sub>F\<close> [0] 1000) is
"\<lambda>f x. if x \<in> dom f then the (f x) else undefined" .
declare [[coercion proj]]
@@ -89,7 +89,7 @@
"Pi' I A = { P. domain P = I \<and> (\<forall>i. i \<in> I \<longrightarrow> (P)\<^sub>F i \<in> A i) } "
syntax
- "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>'' _\<in>_./ _)" 10)
+ "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" (\<open>(3\<Pi>'' _\<in>_./ _)\<close> 10)
syntax_consts
"_Pi'" == Pi'
translations
@@ -636,7 +636,7 @@
"Pi\<^sub>F I M \<equiv> PiF I M"
syntax
- "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" ("(3\<Pi>\<^sub>F _\<in>_./ _)" 10)
+ "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure" (\<open>(3\<Pi>\<^sub>F _\<in>_./ _)\<close> 10)
syntax_consts
"_PiF" == PiF
translations
--- 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 ("\<I>'(_ ; _')") where
+ mutual_information_Pow (\<open>\<I>'(_ ; _')\<close>) where
"\<I>(X ; Y) \<equiv> 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 ("\<H>'(_')") where
+ entropy_Pow (\<open>\<H>'(_')\<close>) where
"\<H>(X) \<equiv> 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 ("\<I>'( _ ; _ | _ ')") where
+ conditional_mutual_information_Pow (\<open>\<I>'( _ ; _ | _ ')\<close>) where
"\<I>(X ; Y | Z) \<equiv> 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)) \<partial>distr M (S \<Otimes>\<^sub>M T) (\<lambda>x. (X x, Y x)))"
abbreviation (in information_space)
- conditional_entropy_Pow ("\<H>'(_ | _')") where
+ conditional_entropy_Pow (\<open>\<H>'(_ | _')\<close>) where
"\<H>(X | Y) \<equiv> conditional_entropy b (count_space (X`space M)) (count_space (Y`space M)) X Y"
lemma (in information_space) conditional_entropy_generic_eq:
--- 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 \<open>Conflicting notation from \<^theory>\<open>HOL-Analysis.Infinite_Sum\<close>\<close>
-no_notation Infinite_Sum.abs_summable_on (infixr "abs'_summable'_on" 46)
+no_notation Infinite_Sum.abs_summable_on (infixr \<open>abs'_summable'_on\<close> 46)
lemma AE_emeasure_singleton:
assumes x: "emeasure M {x} \<noteq> 0" and ae: "AE x in M. P x" shows "P x"
--- 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 \<open>Conflicting notation from \<^theory>\<open>HOL-Analysis.Infinite_Sum\<close>\<close>
-no_notation Infinite_Sum.abs_summable_on (infixr "abs'_summable'_on" 46)
+no_notation Infinite_Sum.abs_summable_on (infixr \<open>abs'_summable'_on\<close> 46)
subsection \<open>Preliminaries\<close>
--- 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 ("\<mu>G") where
+definition mu_G (\<open>\<mu>G\<close>) where
"\<mu>G A = (THE x. \<forall>J\<subseteq>I. finite J \<longrightarrow> (\<forall>X\<in>sets (Pi\<^sub>M J M). A = emb I J X \<longrightarrow> x = emeasure (P J) X))"
definition lim :: "('i \<Rightarrow> 'a) measure" where
--- 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 \<subseteq> I" "finite J" "X \<in> sets (\<Pi>\<^sub>M i\<in>J. borel)"
shows "lim (emb I J X) = P J X"
proof (rule emeasure_lim)
- write mu_G ("\<mu>G")
+ write mu_G (\<open>\<mu>G\<close>)
interpret generator: algebra "space (PiM I (\<lambda>i. borel))" generator
by (rule algebra_generator)
--- 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 \<equiv> rel_pmf (ord_option ord)"
locale ord_spmf_syntax begin
-notation ord_spmf (infix "\<sqsubseteq>\<index>" 60)
+notation ord_spmf (infix \<open>\<sqsubseteq>\<index>\<close> 60)
end
lemma ord_spmf_map_spmf1: "ord_spmf R (map_spmf f p) = ord_spmf (\<lambda>x. R (f x)) p"
@@ -1805,7 +1805,7 @@
subsection \<open>Restrictions on spmfs\<close>
-definition restrict_spmf :: "'a spmf \<Rightarrow> 'a set \<Rightarrow> 'a spmf" (infixl "\<upharpoonleft>" 110)
+definition restrict_spmf :: "'a spmf \<Rightarrow> 'a set \<Rightarrow> 'a spmf" (infixl \<open>\<upharpoonleft>\<close> 110)
where "p \<upharpoonleft> A = map_pmf (\<lambda>x. x \<bind> (\<lambda>y. if y \<in> A then Some y else None)) p"
lemma set_restrict_spmf [simp]: "set_spmf (p \<upharpoonleft> A) = set_spmf p \<inter> A"
@@ -2621,7 +2621,7 @@
subsection \<open>Try\<close>
-definition try_spmf :: "'a spmf \<Rightarrow> 'a spmf \<Rightarrow> 'a spmf" ("TRY _ ELSE _" [0,60] 59)
+definition try_spmf :: "'a spmf \<Rightarrow> 'a spmf \<Rightarrow> 'a spmf" (\<open>TRY _ ELSE _\<close> [0,60] 59)
where "try_spmf p q = bind_pmf p (\<lambda>x. case x of None \<Rightarrow> q | Some y \<Rightarrow> return_spmf y)"
lemma try_spmf_lossless [simp]:
--- 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 ("\<I>'( _ ; _ ')")
+ mutual_information_Pow (\<open>\<I>'( _ ; _ ')\<close>)
notation (in dining_cryptographers_space)
- entropy_Pow ("\<H>'( _ ')")
+ entropy_Pow (\<open>\<H>'( _ ')\<close>)
notation (in dining_cryptographers_space)
- conditional_entropy_Pow ("\<H>'( _ | _ ')")
+ conditional_entropy_Pow (\<open>\<H>'( _ | _ ')\<close>)
theorem (in dining_cryptographers_space)
"\<I>( inversion ; payer ) = 0"
--- 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
"\<mu> \<equiv> point_measure msgs P"
-abbreviation probability ("\<P>'(_') _") where
+abbreviation probability (\<open>\<P>'(_') _\<close>) where
"\<P>(X) x \<equiv> measure \<mu> (X -` x \<inter> msgs)"
-abbreviation joint_probability ("\<P>'(_ ; _') _") where
+abbreviation joint_probability (\<open>\<P>'(_ ; _') _\<close>) where
"\<P>(X ; Y) x \<equiv> \<P>(\<lambda>x. (X x, Y x)) x"
-no_notation disj (infixr "|" 30)
+no_notation disj (infixr \<open>|\<close> 30)
-abbreviation conditional_probability ("\<P>'(_ | _') _") where
+abbreviation conditional_probability (\<open>\<P>'(_ | _') _\<close>) where
"\<P>(X | Y) x \<equiv> (\<P>(X ; Y) x) / \<P>(Y) (snd`x)"
notation
- entropy_Pow ("\<H>'( _ ')")
+ entropy_Pow (\<open>\<H>'( _ ')\<close>)
notation
- conditional_entropy_Pow ("\<H>'( _ | _ ')")
+ conditional_entropy_Pow (\<open>\<H>'( _ | _ ')\<close>)
notation
- mutual_information_Pow ("\<I>'( _ ; _ ')")
+ mutual_information_Pow (\<open>\<I>'( _ ; _ ')\<close>)
lemma t_eq_imp_bij_func:
assumes "t xs = t ys"
--- 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 \<Rightarrow> tm \<Rightarrow> tm" (infixr "and" 999) and
- eq :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixr "eq" 999) and
+ "and" :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixr \<open>and\<close> 999) and
+ eq :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixr \<open>eq\<close> 999) and
- Z :: tm ("Z") and
+ Z :: tm (\<open>Z\<close>) and
S :: "tm \<Rightarrow> tm" and
- plus :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl "+" 65) and
- minus :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl "-" 65) and
- times :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl "*" 70) and
+ plus :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl \<open>+\<close> 65) and
+ minus :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl \<open>-\<close> 65) and
+ times :: "tm \<Rightarrow> tm \<Rightarrow> tm" (infixl \<open>*\<close> 70) and
eval :: "tm \<Rightarrow> tm \<Rightarrow> bool" where
--- 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 \<open>..\<close> 28)
+ Dif :: "[bool, bool] => bool" (infixl \<open>:-\<close> 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 \<open>=>\<close> 27)
translations
--- 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" (\<open>[]\<close>)
+ Cons :: "'a => 'a list => 'a list" (infixr \<open>#\<close> 65)
text \<open>List enumeration\<close>
nonterminal list_args
syntax
- "" :: "'a \<Rightarrow> list_args" ("_")
- "_list_args" :: "'a \<Rightarrow> list_args \<Rightarrow> list_args" ("_,/ _")
- "_list" :: "list_args => 'a list" ("[(_)]")
+ "" :: "'a \<Rightarrow> list_args" (\<open>_\<close>)
+ "_list_args" :: "'a \<Rightarrow> list_args \<Rightarrow> list_args" (\<open>_,/ _\<close>)
+ "_list" :: "list_args => 'a list" (\<open>[(_)]\<close>)
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 (\<open>23\<close>) and
+ nat24 :: nat (\<open>24\<close>) and
+ nat25 :: nat (\<open>25\<close>) and
age :: "[person, nat] => bool" and
--- 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 \<Rightarrow> ty \<Rightarrow> ty" (infixr "->" 20) and
+ arrow :: "ty \<Rightarrow> ty \<Rightarrow> ty" (infixr \<open>->\<close> 20) and
typeof :: "[tm, ty] \<Rightarrow> bool" and
anyterm :: tm
where common_typeof: "
--- 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 "\<rightarrow>\<^sub>\<eta>" 50)
+ eta :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<eta>\<close> 50)
where
eta [simp, intro]: "\<not> free s 0 ==> Abs (s \<degree> Var 0) \<rightarrow>\<^sub>\<eta> s[dummy/0]"
| appL [simp, intro]: "s \<rightarrow>\<^sub>\<eta> t ==> s \<degree> u \<rightarrow>\<^sub>\<eta> t \<degree> u"
@@ -26,11 +26,11 @@
| abs [simp, intro]: "s \<rightarrow>\<^sub>\<eta> t ==> Abs s \<rightarrow>\<^sub>\<eta> Abs t"
abbreviation
- eta_reds :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<eta>\<^sup>*" 50) where
+ eta_reds :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<eta>\<^sup>*\<close> 50) where
"s \<rightarrow>\<^sub>\<eta>\<^sup>* t == eta\<^sup>*\<^sup>* s t"
abbreviation
- eta_red0 :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<eta>\<^sup>=" 50) where
+ eta_red0 :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<eta>\<^sup>=\<close> 50) where
"s \<rightarrow>\<^sub>\<eta>\<^sup>= t == eta\<^sup>=\<^sup>= s t"
inductive_cases eta_cases [elim!]:
--- 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 "\<degree>" 200)
+ | App dB dB (infixl \<open>\<degree>\<close> 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" (\<open>_[_'/_]\<close> [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 \<open>Beta-reduction\<close>
-inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+inductive beta :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<close> 50)
where
beta [simp, intro!]: "Abs s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
@@ -64,7 +64,7 @@
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs s \<rightarrow>\<^sub>\<beta> Abs t"
abbreviation
- beta_reds :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>\<^sup>*" 50) where
+ beta_reds :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<^sup>*\<close> 50) where
"s \<rightarrow>\<^sub>\<beta>\<^sup>* t == beta\<^sup>*\<^sup>* s t"
inductive_cases beta_cases [elim!]:
--- 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 \<open>Environments\<close>
definition
- shift :: "(nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a" ("_\<langle>_:_\<rangle>" [90, 0, 0] 91) where
+ shift :: "(nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a" (\<open>_\<langle>_:_\<rangle>\<close> [90, 0, 0] 91) where
"e\<langle>i:a\<rangle> = (\<lambda>j. if j < i then e j else if j = i then a else e (j - 1))"
lemma shift_eq [simp]: "i = j \<Longrightarrow> (e\<langle>i:T\<rangle>) j = T"
@@ -31,9 +31,9 @@
datatype type =
Atom nat
- | Fun type type (infixr "\<Rightarrow>" 200)
+ | Fun type type (infixr \<open>\<Rightarrow>\<close> 200)
-inductive typing :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile> _ : _" [50, 50, 50] 50)
+inductive typing :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" (\<open>_ \<turnstile> _ : _\<close> [50, 50, 50] 50)
where
Var [intro!]: "env x = T \<Longrightarrow> env \<turnstile> Var x : T"
| Abs [intro!]: "env\<langle>0:T\<rangle> \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs t : (T \<Rightarrow> U)"
@@ -55,11 +55,11 @@
abbreviation
typings_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB list \<Rightarrow> type list \<Rightarrow> bool"
- ("_ \<tturnstile> _ : _" [50, 50, 50] 50) where
+ (\<open>_ \<tturnstile> _ : _\<close> [50, 50, 50] 50) where
"env \<tturnstile> ts : Ts == typings env ts Ts"
abbreviation
- funs :: "type list \<Rightarrow> type \<Rightarrow> type" (infixr "\<Rrightarrow>" 200) where
+ funs :: "type list \<Rightarrow> type \<Rightarrow> type" (infixr \<open>\<Rrightarrow>\<close> 200) where
"Ts \<Rrightarrow> T == foldr Fun Ts T"
--- 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 "\<degree>\<degree>" 150) where
+ list_application :: "dB => dB list => dB" (infixl \<open>\<degree>\<degree>\<close> 150) where
"t \<degree>\<degree> ts == foldl (\<degree>) t ts"
lemma apps_eq_tail_conv [iff]: "(r \<degree>\<degree> ts = s \<degree>\<degree> ts) = (r = s)"
--- 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 @@
\<close>
abbreviation
- list_beta :: "dB list => dB list => bool" (infixl "=>" 50) where
+ list_beta :: "dB list => dB list => bool" (infixl \<open>=>\<close> 50) where
"rs => ss == step1 beta rs ss"
lemma head_Var_reduction:
--- 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 \<open>Parallel reduction\<close>
-inductive par_beta :: "[dB, dB] => bool" (infixl "=>" 50)
+inductive par_beta :: "[dB, dB] => bool" (infixl \<open>=>\<close> 50)
where
var [simp, intro!]: "Var n => Var n"
| abs [simp, intro!]: "s => t ==> Abs s => Abs t"
--- 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 \<Rightarrow> dB \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>s" 50)
- and sredlist :: "dB list \<Rightarrow> dB list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>s]" 50)
+ sred :: "dB \<Rightarrow> dB \<Rightarrow> bool" (infixl \<open>\<rightarrow>\<^sub>s\<close> 50)
+ and sredlist :: "dB list \<Rightarrow> dB list \<Rightarrow> bool" (infixl \<open>[\<rightarrow>\<^sub>s]\<close> 50)
where
"s [\<rightarrow>\<^sub>s] t \<equiv> listrelp (\<rightarrow>\<^sub>s) s t"
| Var: "rs [\<rightarrow>\<^sub>s] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> rs'"
@@ -243,8 +243,8 @@
subsection \<open>Leftmost reduction and weakly normalizing terms\<close>
inductive
- lred :: "dB \<Rightarrow> dB \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>l" 50)
- and lredlist :: "dB list \<Rightarrow> dB list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>l]" 50)
+ lred :: "dB \<Rightarrow> dB \<Rightarrow> bool" (infixl \<open>\<rightarrow>\<^sub>l\<close> 50)
+ and lredlist :: "dB list \<Rightarrow> dB list \<Rightarrow> bool" (infixl \<open>[\<rightarrow>\<^sub>l]\<close> 50)
where
"s [\<rightarrow>\<^sub>l] t \<equiv> listrelp (\<rightarrow>\<^sub>l) s t"
| Var: "rs [\<rightarrow>\<^sub>l] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>l Var x \<degree>\<degree> rs'"
--- 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 @@
\<comment> \<open>A computationally relevant copy of @{term "e \<turnstile> t : T"}\<close>
-inductive rtyping :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ \<turnstile>\<^sub>R _ : _" [50, 50, 50] 50)
+inductive rtyping :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" (\<open>_ \<turnstile>\<^sub>R _ : _\<close> [50, 50, 50] 50)
where
Var: "e x = T \<Longrightarrow> e \<turnstile>\<^sub>R Var x : T"
| Abs: "e\<langle>0:T\<rangle> \<turnstile>\<^sub>R t : U \<Longrightarrow> e \<turnstile>\<^sub>R Abs t : (T \<Rightarrow> U)"
--- 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>\<open>int\<close> is insufficient.\<close>
notation
- less_eq (infix "\<sqsubseteq>" 50) and
- less (infix "\<sqsubset>" 50) and
- top ("\<top>") and
- bot ("\<bottom>") and
- inf (infixl "\<sqinter>" 70) and
- sup (infixl "\<squnion>" 65)
+ less_eq (infix \<open>\<sqsubseteq>\<close> 50) and
+ less (infix \<open>\<sqsubset>\<close> 50) and
+ top (\<open>\<top>\<close>) and
+ bot (\<open>\<bottom>\<close>) and
+ inf (infixl \<open>\<sqinter>\<close> 70) and
+ sup (infixl \<open>\<squnion>\<close> 65)
declare [[quickcheck_narrowing_active = false, quickcheck_timeout = 3600]]
@@ -133,11 +133,11 @@
no_notation
- less_eq (infix "\<sqsubseteq>" 50) and
- less (infix "\<sqsubset>" 50) and
- inf (infixl "\<sqinter>" 70) and
- sup (infixl "\<squnion>" 65) and
- top ("\<top>") and
- bot ("\<bottom>")
+ less_eq (infix \<open>\<sqsubseteq>\<close> 50) and
+ less (infix \<open>\<sqsubset>\<close> 50) and
+ inf (infixl \<open>\<sqinter>\<close> 70) and
+ sup (infixl \<open>\<squnion>\<close> 65) and
+ top (\<open>\<top>\<close>) and
+ bot (\<open>\<bottom>\<close>)
end
--- 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 \<open>Lifted constant definitions\<close>
-lift_definition fnil :: "'a fset" ("{||}") is "[]" parametric list.ctr_transfer(1) .
+lift_definition fnil :: "'a fset" (\<open>{||}\<close>) is "[]" parametric list.ctr_transfer(1) .
lift_definition fcons :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" is Cons parametric list.ctr_transfer(2)
by simp
@@ -97,19 +97,19 @@
nonterminal fset_args
syntax
- "" :: "'a \<Rightarrow> fset_args" ("_")
- "_fset_args" :: "'a \<Rightarrow> fset_args \<Rightarrow> fset_args" ("_,/ _")
- "_fset" :: "fset_args => 'a fset" ("{|(_)|}")
+ "" :: "'a \<Rightarrow> fset_args" (\<open>_\<close>)
+ "_fset_args" :: "'a \<Rightarrow> fset_args \<Rightarrow> fset_args" (\<open>_,/ _\<close>)
+ "_fset" :: "fset_args => 'a fset" (\<open>{|(_)|}\<close>)
syntax_consts
"_fset_args" "_fset" == fcons
translations
"{|x, xs|}" == "CONST fcons x {|xs|}"
"{|x|}" == "CONST fcons x {||}"
-lift_definition fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<in>|" 50) is "\<lambda>x xs. x \<in> set xs"
+lift_definition fmember :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<in>|\<close> 50) is "\<lambda>x xs. x \<in> set xs"
parametric member_transfer by simp
-abbreviation notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50) where
+abbreviation notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<notin>|\<close> 50) where
"x |\<notin>| S \<equiv> \<not> (x |\<in>| S)"
lemma fmember_fmap[simp]: "a |\<in>| fmap f X = (\<exists>b. b |\<in>| X \<and> a = f b)"
--- 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.\<close>
quotient_type
-('a, 'b) fun' (infixr "\<rightarrow>" 55) = "'a \<Rightarrow> 'b" / "(=)"
+('a, 'b) fun' (infixr \<open>\<rightarrow>\<close> 55) = "'a \<Rightarrow> 'b" / "(=)"
by (simp add: identity_equivp)
quotient_definition "comp' :: ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'b) \<rightarrow> 'a \<rightarrow> 'c" is
--- 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 @@
\<close>
definition
- list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
+ list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix \<open>\<approx>\<close> 50)
where
[simp]: "xs \<approx> ys \<longleftrightarrow> set xs = set ys"
@@ -199,7 +199,7 @@
is "Nil :: 'a list" done
abbreviation
- empty_fset ("{||}")
+ empty_fset (\<open>{||}\<close>)
where
"{||} \<equiv> bot :: 'a fset"
@@ -208,7 +208,7 @@
is "sub_list :: ('a list \<Rightarrow> 'a list \<Rightarrow> bool)" by simp
abbreviation
- subset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50)
+ subset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<subseteq>|\<close> 50)
where
"xs |\<subseteq>| ys \<equiv> xs \<le> ys"
@@ -218,7 +218,7 @@
"xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)"
abbreviation
- psubset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50)
+ psubset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<subset>|\<close> 50)
where
"xs |\<subset>| ys \<equiv> xs < ys"
@@ -227,7 +227,7 @@
is "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by simp
abbreviation
- union_fset (infixl "|\<union>|" 65)
+ union_fset (infixl \<open>|\<union>|\<close> 65)
where
"xs |\<union>| ys \<equiv> sup xs (ys::'a fset)"
@@ -236,7 +236,7 @@
is "inter_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" by simp
abbreviation
- inter_fset (infixl "|\<inter>|" 65)
+ inter_fset (infixl \<open>|\<inter>|\<close> 65)
where
"xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)"
@@ -290,9 +290,9 @@
nonterminal fset_args
syntax
- "" :: "'a \<Rightarrow> fset_args" ("_")
- "_fset_args" :: "'a \<Rightarrow> fset_args \<Rightarrow> fset_args" ("_,/ _")
- "_fset" :: "fset_args => 'a fset" ("{|(_)|}")
+ "" :: "'a \<Rightarrow> fset_args" (\<open>_\<close>)
+ "_fset_args" :: "'a \<Rightarrow> fset_args \<Rightarrow> fset_args" (\<open>_,/ _\<close>)
+ "_fset" :: "fset_args => 'a fset" (\<open>{|(_)|}\<close>)
syntax_consts
"_fset_args" "_fset" == insert_fset
translations
@@ -305,12 +305,12 @@
"fset_member :: 'a fset \<Rightarrow> 'a \<Rightarrow> bool" is "List.member" by fastforce
abbreviation
- in_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<in>|" 50)
+ in_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<in>|\<close> 50)
where
"x |\<in>| S \<equiv> fset_member S x"
abbreviation
- notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50)
+ notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix \<open>|\<notin>|\<close> 50)
where
"x |\<notin>| S \<equiv> \<not> (x |\<in>| S)"
@@ -1048,7 +1048,7 @@
and a proof of equivalence\<close>
inductive
- list_eq2 :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<approx>2 _")
+ list_eq2 :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (\<open>_ \<approx>2 _\<close>)
where
"(a # b # xs) \<approx>2 (b # a # xs)"
| "[] \<approx>2 []"
@@ -1156,7 +1156,7 @@
\<close>
no_notation
- list_eq (infix "\<approx>" 50) and
- list_eq2 (infix "\<approx>2" 50)
+ list_eq (infix \<open>\<approx>\<close> 50) and
+ list_eq2 (infix \<open>\<approx>2\<close> 50)
end
--- 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 \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
+ intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix \<open>\<approx>\<close> 50)
where
"intrel (x, y) (u, v) = (x + v = u + y)"
--- 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 \<Rightarrow> freemsg \<Rightarrow> bool" (infixl "\<sim>" 50)
+ msgrel::"freemsg \<Rightarrow> freemsg \<Rightarrow> bool" (infixl \<open>\<sim>\<close> 50)
where
CD: "CRYPT K (DECRYPT K X) \<sim> X"
| DC: "DECRYPT K (CRYPT K X) \<sim> X"
--- 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 \<times> int) \<Rightarrow> (int \<times> int) \<Rightarrow> bool" (infix "\<approx>" 50)
+ ratrel :: "(int \<times> int) \<Rightarrow> (int \<times> int) \<Rightarrow> bool" (infix \<open>\<approx>\<close> 50)
where
[simp]: "x \<approx> y \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
--- 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 \<Rightarrow> real) \<Rightarrow> 'a :: multiseries \<Rightarrow> basis \<Rightarrow> bool"
- (infix "(expands'_to)" 50) where
+ (infix \<open>(expands'_to)\<close> 50) where
"is_expansion F basis \<Longrightarrow> eventually (\<lambda>x. eval F x = f x) at_top \<Longrightarrow> (f expands_to F) basis"
lemma dominant_term_expands_to:
@@ -5375,7 +5375,7 @@
definition expansion_with_remainder_term :: "(real \<Rightarrow> real) \<Rightarrow> (real \<Rightarrow> real) set \<Rightarrow> bool" where
"expansion_with_remainder_term _ _ = True"
-notation (output) expansion_with_remainder_term (infixl "+" 10)
+notation (output) expansion_with_remainder_term (infixl \<open>+\<close> 10)
ML_file \<open>asymptotic_basis.ML\<close>
ML_file \<open>exp_log_expression.ML\<close>
--- 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 \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...*)
nonterminal mtuple_args
syntax
- "" :: "'a \<Rightarrow> mtuple_args" ("_")
- "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" ("_,/ _")
- "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)")
+ "" :: "'a \<Rightarrow> mtuple_args" (\<open>_\<close>)
+ "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" (\<open>_,/ _\<close>)
+ "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" (\<open>(2\<lbrace>_,/ _\<rbrace>)\<close>)
syntax_consts
"_MTuple_args" "_MTuple" \<rightleftharpoons> MPair
translations
--- 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 @@
\<close>
axiomatization
- boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "div'_b" 70) and
- boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "mod'_b" 70)
+ boogie_div :: "int \<Rightarrow> int \<Rightarrow> int" (infixl \<open>div'_b\<close> 70) and
+ boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl \<open>mod'_b\<close> 70)
--- 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} \<Rightarrow> 'a :: {one,zero,ord}) \<Rightarrow> 'real set \<Rightarrow> bool" and
joinpaths :: "('real \<Rightarrow> 'a) \<Rightarrow> ('real \<Rightarrow> 'a) \<Rightarrow> 'real \<Rightarrow> 'a"
begin
-notation piecewise_C1 (infixr "piecewise'_C1'_differentiable'_on" 50)
-notation joinpaths (infixr "+++" 75)
+notation piecewise_C1 (infixr \<open>piecewise'_C1'_differentiable'_on\<close> 50)
+notation joinpaths (infixr \<open>+++\<close> 75)
lemma
\<open>(\<And>v1. \<forall>v0. (rec_join v0 = v1 \<and>
@@ -823,7 +823,7 @@
ground_resolution :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" and
is_least_false_clause :: "'afset \<Rightarrow> 'a \<Rightarrow> bool" and
fset :: "'afset \<Rightarrow> 'a set" and
- union_fset :: "'afset \<Rightarrow> 'afset \<Rightarrow> 'afset" (infixr "|\<union>|" 50)
+ union_fset :: "'afset \<Rightarrow> 'afset \<Rightarrow> 'afset" (infixr \<open>|\<union>|\<close> 50)
begin
term "a |\<union>| b"
--- 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" (\<open>_ :: _\<close> [4, 0] 3)
(*>*)
chapter \<open>HOL-\SPARK{} Reference\<close>
--- 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
\<close>
-definition sdiv :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "sdiv" 70) where
+definition sdiv :: "int \<Rightarrow> int \<Rightarrow> int" (infixl \<open>sdiv\<close> 70) where
"a sdiv b = sgn a * sgn b * (\<bar>a\<bar> div \<bar>b\<bar>)"
lemma sdiv_minus_dividend: "- a sdiv b = - (a sdiv b)"
@@ -55,7 +55,7 @@
"fun_upds f xs y z = (if z \<in> xs then y else f z)"
syntax
- "_updsbind" :: "['a, 'a] => updbind" ("(2_ [:=]/ _)")
+ "_updsbind" :: "['a, 'a] => updbind" (\<open>(2_ [:=]/ _)\<close>)
syntax_consts
"_updsbind" == fun_upds
--- 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 \<Rightarrow> 'b) \<Rightarrow> updbinds \<Rightarrow> ('a \<Rightarrow> 'b)" ("_\<langle>_\<rangle>" [900,0] 900)
+ "_statespace_updates" :: "('a \<Rightarrow> 'b) \<Rightarrow> updbinds \<Rightarrow> ('a \<Rightarrow> 'b)" (\<open>_\<langle>_\<rangle>\<close> [900,0] 900)
(*>*)
--- 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.\<close>
syntax
- "_statespace_lookup" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" ("_\<cdot>_" [60, 60] 60)
+ "_statespace_lookup" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c" (\<open>_\<cdot>_\<close> [60, 60] 60)
"_statespace_update" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'c \<Rightarrow> ('a \<Rightarrow> 'b)"
- "_statespace_updates" :: "('a \<Rightarrow> 'b) \<Rightarrow> updbinds \<Rightarrow> ('a \<Rightarrow> 'b)" ("_<_>" [900, 0] 900)
+ "_statespace_updates" :: "('a \<Rightarrow> 'b) \<Rightarrow> updbinds \<Rightarrow> ('a \<Rightarrow> 'b)" (\<open>_<_>\<close> [900, 0] 900)
translations
"_statespace_updates f (_updbinds b bs)" ==
--- 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 \<Rightarrow> 'a" ("(ACT _)")
+ "_ACT" :: "lift \<Rightarrow> 'a" (\<open>(ACT _)\<close>)
- "_before" :: "lift \<Rightarrow> lift" ("($_)" [100] 99)
- "_after" :: "lift \<Rightarrow> lift" ("(_$)" [100] 99)
- "_unchanged" :: "lift \<Rightarrow> lift" ("(unchanged _)" [100] 99)
+ "_before" :: "lift \<Rightarrow> lift" (\<open>($_)\<close> [100] 99)
+ "_after" :: "lift \<Rightarrow> lift" (\<open>(_$)\<close> [100] 99)
+ "_unchanged" :: "lift \<Rightarrow> lift" (\<open>(unchanged _)\<close> [100] 99)
(*** Priming: same as "after" ***)
- "_prime" :: "lift \<Rightarrow> lift" ("(_`)" [100] 99)
+ "_prime" :: "lift \<Rightarrow> lift" (\<open>(_`)\<close> [100] 99)
- "_Enabled" :: "lift \<Rightarrow> lift" ("(Enabled _)" [100] 100)
+ "_Enabled" :: "lift \<Rightarrow> lift" (\<open>(Enabled _)\<close> [100] 100)
translations
"ACT A" => "(A::state*state \<Rightarrow> _)"
@@ -59,8 +59,8 @@
where angle_def: "AnAct A v \<equiv> ACT (A \<and> \<not> unchanged v)"
syntax
- "_SqAct" :: "[lift, lift] \<Rightarrow> lift" ("([_]'_(_))" [0, 1000] 99)
- "_AnAct" :: "[lift, lift] \<Rightarrow> lift" ("(<_>'_(_))" [0, 1000] 99)
+ "_SqAct" :: "[lift, lift] \<Rightarrow> lift" (\<open>([_]'_(_))\<close> [0, 1000] 99)
+ "_AnAct" :: "[lift, lift] \<Rightarrow> lift" (\<open>(<_>'_(_))\<close> [0, 1000] 99)
translations
"_SqAct" == "CONST SqAct"
"_AnAct" == "CONST AnAct"
--- 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 \<Rightarrow> 'a" ("(TEMP _)")
- "_Init" :: "lift \<Rightarrow> lift" ("(Init _)"[40] 50)
+ "_TEMP" :: "lift \<Rightarrow> 'a" (\<open>(TEMP _)\<close>)
+ "_Init" :: "lift \<Rightarrow> lift" (\<open>(Init _)\<close>[40] 50)
translations
"TEMP F" => "(F::behavior \<Rightarrow> _)"
"_Init" == "CONST Initial"
--- 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 \<equiv> f (x w) (y w) (z w)"
(* "Rigid" quantification (logic level) *)
-definition RAll :: "('a \<Rightarrow> ('w::world) form) \<Rightarrow> 'w form" (binder "Rall " 10)
+definition RAll :: "('a \<Rightarrow> ('w::world) form) \<Rightarrow> 'w form" (binder \<open>Rall \<close> 10)
where unl_Rall: "(Rall x. A x) w \<equiv> \<forall>x. A x w"
-definition REx :: "('a \<Rightarrow> ('w::world) form) \<Rightarrow> 'w form" (binder "Rex " 10)
+definition REx :: "('a \<Rightarrow> ('w::world) form) \<Rightarrow> 'w form" (binder \<open>Rex \<close> 10)
where unl_Rex: "(Rex x. A x) w \<equiv> \<exists>x. A x w"
-definition REx1 :: "('a \<Rightarrow> ('w::world) form) \<Rightarrow> 'w form" (binder "Rex! " 10)
+definition REx1 :: "('a \<Rightarrow> ('w::world) form) \<Rightarrow> 'w form" (binder \<open>Rex! \<close> 10)
where unl_Rex1: "(Rex! x. A x) w \<equiv> \<exists>!x. A x w"
@@ -46,56 +46,56 @@
nonterminal lift and liftargs
syntax
- "" :: "id \<Rightarrow> lift" ("_")
- "" :: "longid \<Rightarrow> lift" ("_")
- "" :: "var \<Rightarrow> lift" ("_")
- "_applC" :: "[lift, cargs] \<Rightarrow> lift" ("(1_/ _)" [1000, 1000] 999)
- "" :: "lift \<Rightarrow> lift" ("'(_')")
- "_lambda" :: "[idts, 'a] \<Rightarrow> lift" ("(3\<lambda>_./ _)" [0, 3] 3)
- "_constrain" :: "[lift, type] \<Rightarrow> lift" ("(_::_)" [4, 0] 3)
- "" :: "lift \<Rightarrow> liftargs" ("_")
- "_liftargs" :: "[lift, liftargs] \<Rightarrow> liftargs" ("_,/ _")
- "_Valid" :: "lift \<Rightarrow> bool" ("(\<turnstile> _)" 5)
- "_holdsAt" :: "['a, lift] \<Rightarrow> bool" ("(_ \<Turnstile> _)" [100,10] 10)
+ "" :: "id \<Rightarrow> lift" (\<open>_\<close>)
+ "" :: "longid \<Rightarrow> lift" (\<open>_\<close>)
+ "" :: "var \<Rightarrow> lift" (\<open>_\<close>)
+ "_applC" :: "[lift, cargs] \<Rightarrow> lift" (\<open>(1_/ _)\<close> [1000, 1000] 999)
+ "" :: "lift \<Rightarrow> lift" (\<open>'(_')\<close>)
+ "_lambda" :: "[idts, 'a] \<Rightarrow> lift" (\<open>(3\<lambda>_./ _)\<close> [0, 3] 3)
+ "_constrain" :: "[lift, type] \<Rightarrow> lift" (\<open>(_::_)\<close> [4, 0] 3)
+ "" :: "lift \<Rightarrow> liftargs" (\<open>_\<close>)
+ "_liftargs" :: "[lift, liftargs] \<Rightarrow> liftargs" (\<open>_,/ _\<close>)
+ "_Valid" :: "lift \<Rightarrow> bool" (\<open>(\<turnstile> _)\<close> 5)
+ "_holdsAt" :: "['a, lift] \<Rightarrow> bool" (\<open>(_ \<Turnstile> _)\<close> [100,10] 10)
(* Syntax for lifted expressions outside the scope of \<turnstile> or |= *)
- "_LIFT" :: "lift \<Rightarrow> 'a" ("LIFT _")
+ "_LIFT" :: "lift \<Rightarrow> 'a" (\<open>LIFT _\<close>)
(* generic syntax for lifted constants and functions *)
- "_const" :: "'a \<Rightarrow> lift" ("(#_)" [1000] 999)
- "_lift" :: "['a, lift] \<Rightarrow> lift" ("(_<_>)" [1000] 999)
- "_lift2" :: "['a, lift, lift] \<Rightarrow> lift" ("(_<_,/ _>)" [1000] 999)
- "_lift3" :: "['a, lift, lift, lift] \<Rightarrow> lift" ("(_<_,/ _,/ _>)" [1000] 999)
+ "_const" :: "'a \<Rightarrow> lift" (\<open>(#_)\<close> [1000] 999)
+ "_lift" :: "['a, lift] \<Rightarrow> lift" (\<open>(_<_>)\<close> [1000] 999)
+ "_lift2" :: "['a, lift, lift] \<Rightarrow> lift" (\<open>(_<_,/ _>)\<close> [1000] 999)
+ "_lift3" :: "['a, lift, lift, lift] \<Rightarrow> lift" (\<open>(_<_,/ _,/ _>)\<close> [1000] 999)
(* concrete syntax for common infix functions: reuse same symbol *)
- "_liftEqu" :: "[lift, lift] \<Rightarrow> lift" ("(_ =/ _)" [50,51] 50)
- "_liftNeq" :: "[lift, lift] \<Rightarrow> lift" ("(_ \<noteq>/ _)" [50,51] 50)
- "_liftNot" :: "lift \<Rightarrow> lift" ("(\<not> _)" [40] 40)
- "_liftAnd" :: "[lift, lift] \<Rightarrow> lift" ("(_ \<and>/ _)" [36,35] 35)
- "_liftOr" :: "[lift, lift] \<Rightarrow> lift" ("(_ \<or>/ _)" [31,30] 30)
- "_liftImp" :: "[lift, lift] \<Rightarrow> lift" ("(_ \<longrightarrow>/ _)" [26,25] 25)
- "_liftIf" :: "[lift, lift, lift] \<Rightarrow> lift" ("(if (_)/ then (_)/ else (_))" 10)
- "_liftPlus" :: "[lift, lift] \<Rightarrow> lift" ("(_ +/ _)" [66,65] 65)
- "_liftMinus" :: "[lift, lift] \<Rightarrow> lift" ("(_ -/ _)" [66,65] 65)
- "_liftTimes" :: "[lift, lift] \<Rightarrow> lift" ("(_ */ _)" [71,70] 70)
- "_liftDiv" :: "[lift, lift] \<Rightarrow> lift" ("(_ div _)" [71,70] 70)
- "_liftMod" :: "[lift, lift] \<Rightarrow> lift" ("(_ mod _)" [71,70] 70)
- "_liftLess" :: "[lift, lift] \<Rightarrow> lift" ("(_/ < _)" [50, 51] 50)
- "_liftLeq" :: "[lift, lift] \<Rightarrow> lift" ("(_/ \<le> _)" [50, 51] 50)
- "_liftMem" :: "[lift, lift] \<Rightarrow> lift" ("(_/ \<in> _)" [50, 51] 50)
- "_liftNotMem" :: "[lift, lift] \<Rightarrow> lift" ("(_/ \<notin> _)" [50, 51] 50)
- "_liftFinset" :: "liftargs \<Rightarrow> lift" ("{(_)}")
+ "_liftEqu" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ =/ _)\<close> [50,51] 50)
+ "_liftNeq" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ \<noteq>/ _)\<close> [50,51] 50)
+ "_liftNot" :: "lift \<Rightarrow> lift" (\<open>(\<not> _)\<close> [40] 40)
+ "_liftAnd" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ \<and>/ _)\<close> [36,35] 35)
+ "_liftOr" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ \<or>/ _)\<close> [31,30] 30)
+ "_liftImp" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ \<longrightarrow>/ _)\<close> [26,25] 25)
+ "_liftIf" :: "[lift, lift, lift] \<Rightarrow> lift" (\<open>(if (_)/ then (_)/ else (_))\<close> 10)
+ "_liftPlus" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ +/ _)\<close> [66,65] 65)
+ "_liftMinus" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ -/ _)\<close> [66,65] 65)
+ "_liftTimes" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ */ _)\<close> [71,70] 70)
+ "_liftDiv" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ div _)\<close> [71,70] 70)
+ "_liftMod" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ mod _)\<close> [71,70] 70)
+ "_liftLess" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_/ < _)\<close> [50, 51] 50)
+ "_liftLeq" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_/ \<le> _)\<close> [50, 51] 50)
+ "_liftMem" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_/ \<in> _)\<close> [50, 51] 50)
+ "_liftNotMem" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_/ \<notin> _)\<close> [50, 51] 50)
+ "_liftFinset" :: "liftargs \<Rightarrow> lift" (\<open>{(_)}\<close>)
(** TODO: syntax for lifted collection / comprehension **)
- "_liftPair" :: "[lift,liftargs] \<Rightarrow> lift" ("(1'(_,/ _'))")
+ "_liftPair" :: "[lift,liftargs] \<Rightarrow> lift" (\<open>(1'(_,/ _'))\<close>)
(* infix syntax for list operations *)
- "_liftCons" :: "[lift, lift] \<Rightarrow> lift" ("(_ #/ _)" [65,66] 65)
- "_liftApp" :: "[lift, lift] \<Rightarrow> lift" ("(_ @/ _)" [65,66] 65)
- "_liftList" :: "liftargs \<Rightarrow> lift" ("[(_)]")
+ "_liftCons" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ #/ _)\<close> [65,66] 65)
+ "_liftApp" :: "[lift, lift] \<Rightarrow> lift" (\<open>(_ @/ _)\<close> [65,66] 65)
+ "_liftList" :: "liftargs \<Rightarrow> lift" (\<open>[(_)]\<close>)
(* Rigid quantification (syntax level) *)
- "_RAll" :: "[idts, lift] \<Rightarrow> lift" ("(3\<forall>_./ _)" [0, 10] 10)
- "_REx" :: "[idts, lift] \<Rightarrow> lift" ("(3\<exists>_./ _)" [0, 10] 10)
- "_REx1" :: "[idts, lift] \<Rightarrow> lift" ("(3\<exists>!_./ _)" [0, 10] 10)
+ "_RAll" :: "[idts, lift] \<Rightarrow> lift" (\<open>(3\<forall>_./ _)\<close> [0, 10] 10)
+ "_REx" :: "[idts, lift] \<Rightarrow> lift" (\<open>(3\<exists>_./ _)\<close> [0, 10] 10)
+ "_REx1" :: "[idts, lift] \<Rightarrow> lift" (\<open>(3\<exists>!_./ _)\<close> [0, 10] 10)
translations
"_const" == "CONST const"
--- 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 \<equiv> x s i"
syntax
- "_slice" :: "[lift, 'a] \<Rightarrow> lift" ("(_!_)" [70,70] 70)
+ "_slice" :: "[lift, 'a] \<Rightarrow> lift" (\<open>(_!_)\<close> [70,70] 70)
translations
"_slice" \<rightleftharpoons> "CONST slice"
@@ -67,8 +67,8 @@
\<and> (res<ch!p>$ = $v)"
syntax
- "_Call" :: "['a, 'b, lift] \<Rightarrow> lift" ("(Call _ _ _)" [90,90,90] 90)
- "_Return" :: "['a, 'b, lift] \<Rightarrow> lift" ("(Return _ _ _)" [90,90,90] 90)
+ "_Call" :: "['a, 'b, lift] \<Rightarrow> lift" (\<open>(Call _ _ _)\<close> [90,90,90] 90)
+ "_Return" :: "['a, 'b, lift] \<Rightarrow> lift" (\<open>(Return _ _ _)\<close> [90,90,90] 90)
translations
"_Call" \<rightleftharpoons> "CONST ACall"
"_Return" \<rightleftharpoons> "CONST AReturn"
--- 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 \<Rightarrow> bool"
syntax
- "_PRED" :: "lift \<Rightarrow> 'a" ("PRED _")
- "_stvars" :: "lift \<Rightarrow> bool" ("basevars _")
+ "_PRED" :: "lift \<Rightarrow> 'a" (\<open>PRED _\<close>)
+ "_stvars" :: "lift \<Rightarrow> bool" (\<open>basevars _\<close>)
translations
"PRED P" => "(P::state \<Rightarrow> _)"
--- 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] \<Rightarrow> temporal"
(* Quantification over (flexible) state variables *)
- EEx :: "('a stfun \<Rightarrow> temporal) \<Rightarrow> temporal" (binder "Eex " 10)
- AAll :: "('a stfun \<Rightarrow> temporal) \<Rightarrow> temporal" (binder "Aall " 10)
+ EEx :: "('a stfun \<Rightarrow> temporal) \<Rightarrow> temporal" (binder \<open>Eex \<close> 10)
+ AAll :: "('a stfun \<Rightarrow> temporal) \<Rightarrow> temporal" (binder \<open>Aall \<close> 10)
(** concrete syntax **)
syntax
- "_Box" :: "lift \<Rightarrow> lift" ("(\<box>_)" [40] 40)
- "_Dmd" :: "lift \<Rightarrow> lift" ("(\<diamond>_)" [40] 40)
- "_leadsto" :: "[lift,lift] \<Rightarrow> lift" ("(_ \<leadsto> _)" [23,22] 22)
- "_stable" :: "lift \<Rightarrow> lift" ("(stable/ _)")
- "_WF" :: "[lift,lift] \<Rightarrow> lift" ("(WF'(_')'_(_))" [0,60] 55)
- "_SF" :: "[lift,lift] \<Rightarrow> lift" ("(SF'(_')'_(_))" [0,60] 55)
+ "_Box" :: "lift \<Rightarrow> lift" (\<open>(\<box>_)\<close> [40] 40)
+ "_Dmd" :: "lift \<Rightarrow> lift" (\<open>(\<diamond>_)\<close> [40] 40)
+ "_leadsto" :: "[lift,lift] \<Rightarrow> lift" (\<open>(_ \<leadsto> _)\<close> [23,22] 22)
+ "_stable" :: "lift \<Rightarrow> lift" (\<open>(stable/ _)\<close>)
+ "_WF" :: "[lift,lift] \<Rightarrow> lift" (\<open>(WF'(_')'_(_))\<close> [0,60] 55)
+ "_SF" :: "[lift,lift] \<Rightarrow> lift" (\<open>(SF'(_')'_(_))\<close> [0,60] 55)
- "_EEx" :: "[idts, lift] \<Rightarrow> lift" ("(3\<exists>\<exists> _./ _)" [0,10] 10)
- "_AAll" :: "[idts, lift] \<Rightarrow> lift" ("(3\<forall>\<forall> _./ _)" [0,10] 10)
+ "_EEx" :: "[idts, lift] \<Rightarrow> lift" (\<open>(3\<exists>\<exists> _./ _)\<close> [0,10] 10)
+ "_AAll" :: "[idts, lift] \<Rightarrow> lift" (\<open>(3\<forall>\<forall> _./ _)\<close> [0,10] 10)
translations
"_Box" == "CONST Box"
--- 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 \<open>Definitions \<^emph>\<open>on\<close> carrier set\<close>
locale module_on =
- fixes S and scale :: "'a::comm_ring_1 \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr "*s" 75)
+ fixes S and scale :: "'a::comm_ring_1 \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr \<open>*s\<close> 75)
assumes scale_right_distrib_on [algebra_simps]: "x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> a *s (x + y) = a *s x + a *s y"
and scale_left_distrib_on [algebra_simps]: "x \<in> S \<Longrightarrow> (a + b) *s x = a *s x + b *s x"
and scale_scale_on [simp]: "x \<in> S \<Longrightarrow> 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 \<Rightarrow> 'b \<Rightarrow> 'b" (infixr "*a" 75)
- and s2 :: "'a::comm_ring_1 \<Rightarrow> 'c \<Rightarrow> 'c" (infixr "*b" 75) +
+ and s1 :: "'a::comm_ring_1 \<Rightarrow> 'b \<Rightarrow> 'b" (infixr \<open>*a\<close> 75)
+ and s2 :: "'a::comm_ring_1 \<Rightarrow> 'c \<Rightarrow> 'c" (infixr \<open>*b\<close> 75) +
fixes f :: "'b \<Rightarrow> 'c"
assumes add: "\<And>b1 b2. b1 \<in> S1 \<Longrightarrow> b2 \<in> S1 \<Longrightarrow> f (b1 + b2) = f b1 + f b2"
and scale: "\<And>b. b \<in> S1 \<Longrightarrow> 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 \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr "*s" 75)
+ for S and scale :: "'a::field \<Rightarrow> 'b::ab_group_add \<Rightarrow> 'b" (infixr \<open>*s\<close> 75)
begin
definition dim :: "'b set \<Rightarrow> nat"
--- 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 \<open>component'_of\<close> 50)
where "F component_of H == \<exists>G. F ok G & F\<squnion>G = H"
-definition strict_component_of :: "'a program\<Rightarrow>'a program=> bool" (infixl "strict'_component'_of" 50)
+definition strict_component_of :: "'a program\<Rightarrow>'a program=> bool" (infixl \<open>strict'_component'_of\<close> 50)
where "F strict_component_of H == F component_of H & F\<noteq>H"
definition preserves :: "('a=>'b) => 'a program set"
--- 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 \<in> Acts F; s \<in> reachable F; (s,s') \<in> act |]
==> s' \<in> 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 \<open>Co\<close> 60) where
"A Co B == {F. F \<in> (reachable F \<inter> 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 \<open>Unless\<close> 60) where
"A Unless B == (A-B) Co (A \<union> B)"
definition Stable :: "'a set => 'a program set" where
--- 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 \<open>Detects\<close> 60)
where "A Detects B = (Always (-A \<union> B)) \<inter> (B LeadsTo A)"
-definition Equality :: "['a set, 'a set] => 'a set" (infixl "<==>" 60)
+definition Equality :: "['a set, 'a set] => 'a set" (infixl \<open><==>\<close> 60)
where "A <==> B = (-A \<union> B) \<inter> (A \<union> -B)"
--- 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)
+ (\<open>(3_/ leadsTo[_]/ _)\<close> [80,0,80] 80)
where "leadsETo A CC B = {F. (A,B) \<in> elt CC F}"
definition
LeadsETo :: "['a set, 'a set set, 'a set] => 'a program set"
- ("(3_/ LeadsTo[_]/ _)" [80,0,80] 80)
+ (\<open>(3_/ LeadsTo[_]/ _)\<close> [80,0,80] 80)
where "LeadsETo A CC B =
{F. F \<in> (reachable F Int A) leadsTo[(%C. reachable F Int C) ` CC] B}"
--- 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 \<open>Fols\<close> 65) where
"f Fols g == Increasing g \<inter> Increasing f Int
Always {s. f s \<le> g s} Int
(\<Inter>k. {s. k \<le> g s} LeadsTo {s. k \<le> f s})"
--- 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\<open>Guarantees properties\<close>
-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 \<open>guarantees\<close> 55) where
(*higher than membership, lower than Co*)
"X guarantees Y == {F. \<forall>G. F ok G --> F\<squnion>G \<in> X --> F\<squnion>G \<in> Y}"
@@ -56,13 +56,13 @@
"welldef == {F. Init F \<noteq> {}}"
definition refines :: "['a program, 'a program, 'a program set] => bool"
- ("(3_ refines _ wrt _)" [10,10,10] 10) where
+ (\<open>(3_ refines _ wrt _)\<close> [10,10,10] 10) where
"G refines F wrt X ==
\<forall>H. (F ok H & G ok H & F\<squnion>H \<in> welldef \<inter> X) -->
(G\<squnion>H \<in> welldef \<inter> X)"
definition iso_refines :: "['a program, 'a program, 'a program set] => bool"
- ("(3_ iso'_refines _ wrt _)" [10,10,10] 10) where
+ (\<open>(3_ iso'_refines _ wrt _)\<close> [10,10,10] 10) where
"G iso_refines F wrt X ==
F \<in> welldef \<inter> X --> G \<in> welldef \<inter> X"
--- 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 \<open>pfixLe\<close> 50) where
"xs pfixLe ys == (xs,ys) \<in> genPrefix Le"
abbreviation
- pfixGe :: "[nat list, nat list] => bool" (infixl "pfixGe" 50) where
+ pfixGe :: "[nat list, nat list] => bool" (infixl \<open>pfixGe\<close> 50) where
"xs pfixGe ys == (xs,ys) \<in> genPrefix Ge"
--- 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 == \<Squnion>i \<in> 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" (\<open>(3plam _:_./ _)\<close> 10)
syntax_consts
"_PLam" == PLam
translations
--- 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 \<open>Ensures\<close> 60) where
"A Ensures B == {F. F \<in> (reachable F \<inter> 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 \<open>LeadsTo\<close> 60) where
"A LeadsTo B == {F. F \<in> (reachable F \<inter> A) leadsTo B}"
-notation LeadsTo (infixl "\<longmapsto>w" 60)
+notation LeadsTo (infixl \<open>\<longmapsto>w\<close> 60)
text\<open>Resembles the previous definition of LeadsTo\<close>
--- 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 \<open>co\<close> 60) where
"A co B == {F. \<forall>act \<in> Acts F. act``A \<subseteq> B}"
-definition unless :: "['a set, 'a set] => 'a program set" (infixl "unless" 60) where
+definition unless :: "['a set, 'a set] => 'a program set" (infixl \<open>unless\<close> 60) where
"A unless B == (A-B) co (A \<union> B)"
definition mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set)
--- 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 \<inter> Init G \<noteq> {} *)
definition
- ok :: "['a program, 'a program] => bool" (infixl "ok" 65)
+ ok :: "['a program, 'a program] => bool" (infixl \<open>ok\<close> 65)
where "F ok G == Acts F \<subseteq> AllowedActs G &
Acts G \<subseteq> AllowedActs F"
@@ -26,11 +26,11 @@
\<Inter>i \<in> I. AllowedActs (F i))"
definition
- Join :: "['a program, 'a program] => 'a program" (infixl "\<squnion>" 65)
+ Join :: "['a program, 'a program] => 'a program" (infixl \<open>\<squnion>\<close> 65)
where "F \<squnion> G = mk_program (Init F \<inter> Init G, Acts F \<union> Acts G,
AllowedActs F \<inter> AllowedActs G)"
-definition SKIP :: "'a program" ("\<bottom>")
+definition SKIP :: "'a program" (\<open>\<bottom>\<close>)
where "\<bottom> = mk_program (UNIV, {}, UNIV)"
(*Characterizes safety properties. Used with specifying Allowed*)
@@ -39,8 +39,8 @@
where "safety_prop X \<longleftrightarrow> SKIP \<in> X \<and> (\<forall>G. Acts G \<subseteq> \<Union>(Acts ` X) \<longrightarrow> G \<in> X)"
syntax
- "_JOIN1" :: "[pttrns, 'b set] => 'b set" ("(3\<Squnion>_./ _)" 10)
- "_JOIN" :: "[pttrn, 'a set, 'b set] => 'b set" ("(3\<Squnion>_\<in>_./ _)" 10)
+ "_JOIN1" :: "[pttrns, 'b set] => 'b set" (\<open>(3\<Squnion>_./ _)\<close> 10)
+ "_JOIN" :: "[pttrn, 'a set, 'b set] => 'b set" (\<open>(3\<Squnion>_\<in>_./ _)\<close> 10)
syntax_consts
"_JOIN1" "_JOIN" == JOIN
translations
--- 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. \<exists>act\<in>Acts F. act``A \<subseteq> -A}"
definition
- ensures :: "['a set, 'a set] => 'a program set" (infixl "ensures" 60) where
+ ensures :: "['a set, 'a set] => 'a program set" (infixl \<open>ensures\<close> 60) where
"A ensures B == (A-B co A \<union> B) \<inter> transient (A-B)"
@@ -59,7 +59,7 @@
| Union: "\<forall>A \<in> S. (A,B) \<in> leads F ==> (Union S, B) \<in> 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 \<open>leadsTo\<close> 60) where
\<comment> \<open>visible version of the LEADS-TO relation\<close>
"A leadsTo B == {F. (A,B) \<in> leads F}"
@@ -67,7 +67,7 @@
\<comment> \<open>predicate transformer: the largest set that leads to \<^term>\<open>B\<close>\<close>
"wlt F B == \<Union>{A. F \<in> A leadsTo B}"
-notation leadsTo (infixl "\<longmapsto>" 60)
+notation leadsTo (infixl \<open>\<longmapsto>\<close> 60)
subsection\<open>transient\<close>
--- 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 @@
\<close>
inductive transition :: "file \<Rightarrow> operation \<Rightarrow> file \<Rightarrow> bool"
- ("_ \<midarrow>_\<rightarrow> _" [90, 1000, 90] 100)
+ (\<open>_ \<midarrow>_\<rightarrow> _\<close> [90, 1000, 90] 100)
where
read:
"root \<midarrow>(Read uid text path)\<rightarrow> root"
@@ -475,7 +475,7 @@
running processes over a finite amount of time.
\<close>
-inductive transitions :: "file \<Rightarrow> operation list \<Rightarrow> file \<Rightarrow> bool" ("_ \<Midarrow>_\<Rightarrow> _" [90, 1000, 90] 100)
+inductive transitions :: "file \<Rightarrow> operation list \<Rightarrow> file \<Rightarrow> bool" (\<open>_ \<Midarrow>_\<Rightarrow> _\<close> [90, 1000, 90] 100)
where
nil: "root \<Midarrow>[]\<Rightarrow> root"
| cons: "root \<Midarrow>(x # xs)\<Rightarrow> root''" if "root \<midarrow>x\<rightarrow> root'" and "root' \<Midarrow>xs\<Rightarrow> root''"
--- 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 \<Rightarrow> ZF" where
"Field A == union (Domain A) (Range A)"
-definition app :: "ZF \<Rightarrow> ZF => ZF" (infixl "\<acute>" 90) \<comment> \<open>function application\<close> where
+definition app :: "ZF \<Rightarrow> ZF => ZF" (infixl \<open>\<acute>\<close> 90) \<comment> \<open>function application\<close> where
"f \<acute> x == (THE y. Elem (Opair x y) f)"
definition isFun :: "ZF \<Rightarrow> bool" where
--- 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 \<open>A simple example on quote / antiquote in higher-order abstract syntax.\<close>
-definition var :: "'a \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> nat" ("VAR _" [1000] 999)
+definition var :: "'a \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> nat" (\<open>VAR _\<close> [1000] 999)
where "var x env = env x"
definition Expr :: "(('a \<Rightarrow> nat) \<Rightarrow> nat) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> nat"
where "Expr exp env = exp env"
syntax
- "_Expr" :: "'a \<Rightarrow> 'a" ("EXPR _" [1000] 999)
+ "_Expr" :: "'a \<Rightarrow> 'a" (\<open>EXPR _\<close> [1000] 999)
parse_translation
\<open>[Syntax_Trans.quote_antiquote_tr
--- 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 \<open>Definitions\<close>
-definition bigo :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) set" ("(1O'(_'))")
+definition bigo :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) set" (\<open>(1O'(_'))\<close>)
where "O(f:: 'a \<Rightarrow> 'b) = {h. \<exists>c. \<forall>x. \<bar>h x\<bar> \<le> c * \<bar>f x\<bar>}"
lemma bigo_pos_const:
@@ -407,7 +407,7 @@
subsection \<open>Less than or equal to\<close>
-definition lesso :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" (infixl "<o" 70)
+definition lesso :: "('a \<Rightarrow> 'b::linordered_idom) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" (infixl \<open><o\<close> 70)
where "f <o g = (\<lambda>x. max (f x - g x) 0)"
lemma bigo_lesseq1: "f =o O(h) \<Longrightarrow> \<forall>x. \<bar>g x\<bar> \<le> \<bar>f x\<bar> \<Longrightarrow> g =o O(h)"
--- 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 \<Rightarrow> 'a ctl \<Rightarrow> 'a ctl" (infixr "\<rightarrow>" 75)
+definition imp :: "'a ctl \<Rightarrow> 'a ctl \<Rightarrow> 'a ctl" (infixr \<open>\<rightarrow>\<close> 75)
where "p \<rightarrow> q = - p \<union> q"
lemma [intro!]: "p \<inter> p \<rightarrow> q \<subseteq> q" unfolding imp_def by auto
@@ -53,13 +53,13 @@
\<^cite>\<open>"McMillan-PhDThesis"\<close>.
\<close>
-definition EX ("\<^bold>E\<^bold>X _" [80] 90)
+definition EX (\<open>\<^bold>E\<^bold>X _\<close> [80] 90)
where [simp]: "\<^bold>E\<^bold>X p = {s. \<exists>s'. (s, s') \<in> \<M> \<and> s' \<in> p}"
-definition EF ("\<^bold>E\<^bold>F _" [80] 90)
+definition EF (\<open>\<^bold>E\<^bold>F _\<close> [80] 90)
where [simp]: "\<^bold>E\<^bold>F p = lfp (\<lambda>s. p \<union> \<^bold>E\<^bold>X s)"
-definition EG ("\<^bold>E\<^bold>G _" [80] 90)
+definition EG (\<open>\<^bold>E\<^bold>G _\<close> [80] 90)
where [simp]: "\<^bold>E\<^bold>G p = gfp (\<lambda>s. p \<inter> \<^bold>E\<^bold>X s)"
text \<open>
@@ -67,11 +67,11 @@
\<open>\<^bold>E\<^bold>F\<close> and \<open>\<^bold>E\<^bold>G\<close>.
\<close>
-definition AX ("\<^bold>A\<^bold>X _" [80] 90)
+definition AX (\<open>\<^bold>A\<^bold>X _\<close> [80] 90)
where [simp]: "\<^bold>A\<^bold>X p = - \<^bold>E\<^bold>X - p"
-definition AF ("\<^bold>A\<^bold>F _" [80] 90)
+definition AF (\<open>\<^bold>A\<^bold>F _\<close> [80] 90)
where [simp]: "\<^bold>A\<^bold>F p = - \<^bold>E\<^bold>G - p"
-definition AG ("\<^bold>A\<^bold>G _" [80] 90)
+definition AG (\<open>\<^bold>A\<^bold>G _\<close> [80] 90)
where [simp]: "\<^bold>A\<^bold>G p = - \<^bold>E\<^bold>F - p"
--- 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;
\<close>
-syntax "_cartouche_string" :: \<open>cartouche_position \<Rightarrow> string\<close> ("_")
+syntax "_cartouche_string" :: \<open>cartouche_position \<Rightarrow> string\<close> (\<open>_\<close>)
parse_translation \<open>
[(\<^syntax_const>\<open>_cartouche_string\<close>,
@@ -95,7 +95,7 @@
subsubsection \<open>Nested quotes\<close>
-syntax "_string_string" :: \<open>string_position \<Rightarrow> string\<close> ("_")
+syntax "_string_string" :: \<open>string_position \<Rightarrow> string\<close> (\<open>_\<close>)
parse_translation \<open>
[(\<^syntax_const>\<open>_string_string\<close>, K (string_tr Lexicon.explode_string))]
--- 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 @@
\<close>
datatype shuzi =
- One ("一")
- | Two ("二")
- | Three ("三")
- | Four ("四")
- | Five ("五")
- | Six ("六")
- | Seven ("七")
- | Eight ("八")
- | Nine ("九")
- | Ten ("十")
+ One (\<open>一\<close>)
+ | Two (\<open>二\<close>)
+ | Three (\<open>三\<close>)
+ | Four (\<open>四\<close>)
+ | Five (\<open>五\<close>)
+ | Six (\<open>六\<close>)
+ | Seven (\<open>七\<close>)
+ | Eight (\<open>八\<close>)
+ | Nine (\<open>九\<close>)
+ | Ten (\<open>十\<close>)
-primrec translate :: "shuzi \<Rightarrow> nat" ("|_|" [100] 100) where
+primrec translate :: "shuzi \<Rightarrow> nat" (\<open>|_|\<close> [100] 100) where
"|一| = 1"
|"|二| = 2"
|"|三| = 3"
--- 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 \<open>Lazy evaluation for streams\<close>
codatatype 'a stream =
- SCons (shd: 'a) (stl: "'a stream") (infixr "##" 65)
+ SCons (shd: 'a) (stl: "'a stream") (infixr \<open>##\<close> 65)
primcorec up :: "nat \<Rightarrow> nat stream" where
"up n = n ## up (n + 1)"
@@ -36,14 +36,14 @@
text \<open>Lazy types need not be infinite. We can also have lazy types that are finite.\<close>
datatype 'a llist
- = LNil ("\<^bold>\<lbrakk>\<^bold>\<rbrakk>")
- | LCons (lhd: 'a) (ltl: "'a llist") (infixr "###" 65)
+ = LNil (\<open>\<^bold>\<lbrakk>\<^bold>\<rbrakk>\<close>)
+ | LCons (lhd: 'a) (ltl: "'a llist") (infixr \<open>###\<close> 65)
nonterminal llist_args
syntax
- "" :: "'a \<Rightarrow> llist_args" ("_")
- "_llist_args" :: "'a \<Rightarrow> llist_args \<Rightarrow> llist_args" ("_,/ _")
- "_llist" :: "llist_args => 'a list" ("\<^bold>\<lbrakk>(_)\<^bold>\<rbrakk>")
+ "" :: "'a \<Rightarrow> llist_args" (\<open>_\<close>)
+ "_llist_args" :: "'a \<Rightarrow> llist_args \<Rightarrow> llist_args" (\<open>_,/ _\<close>)
+ "_llist" :: "llist_args => 'a list" (\<open>\<^bold>\<lbrakk>(_)\<^bold>\<rbrakk>\<close>)
syntax_consts
"_llist_args" "_llist" == LCons
translations
@@ -81,7 +81,7 @@
A conversion function to lazy lists is enough.\<close>
primrec lappend :: "'a llist \<Rightarrow> 'a llist \<Rightarrow> 'a llist"
- (infixr "@@" 65) where
+ (infixr \<open>@@\<close> 65) where
"\<^bold>\<lbrakk>\<^bold>\<rbrakk> @@ ys = ys"
| "(x ### xs) @@ ys = x ### (xs @@ ys)"
@@ -113,10 +113,10 @@
section \<open>Branching datatypes\<close>
datatype tree
- = L ("\<spadesuit>")
- | Node tree tree (infix "\<triangle>" 900)
+ = L (\<open>\<spadesuit>\<close>)
+ | Node tree tree (infix \<open>\<triangle>\<close> 900)
-notation (output) Node ("\<triangle>(//\<^bold>l: _//\<^bold>r: _)")
+notation (output) Node (\<open>\<triangle>(//\<^bold>l: _//\<^bold>r: _)\<close>)
code_lazy_type tree
--- 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 \<open>The \<open>\<lesssim>\<close> relation\<close>
-definition less_eq_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<lesssim>" 50)
+definition less_eq_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix \<open>\<lesssim>\<close> 50)
where
"f \<lesssim> g \<longleftrightarrow> (\<exists>c>0. \<exists>n. \<forall>m>n. f m \<le> c * g m)"
@@ -73,7 +73,7 @@
subsection \<open>The \<open>\<approx>\<close> relation, the equivalence relation induced by \<open>\<lesssim>\<close>\<close>
-definition equiv_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<cong>" 50)
+definition equiv_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix \<open>\<cong>\<close> 50)
where
"f \<cong> g \<longleftrightarrow>
(\<exists>c\<^sub>1>0. \<exists>c\<^sub>2>0. \<exists>n. \<forall>m>n. f m \<le> c\<^sub>1 * g m \<and> g m \<le> c\<^sub>2 * f m)"
@@ -110,7 +110,7 @@
subsection \<open>The \<open>\<prec>\<close> relation, the strict part of \<open>\<lesssim>\<close>\<close>
-definition less_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix "\<prec>" 50)
+definition less_fun :: "(nat \<Rightarrow> nat) \<Rightarrow> (nat \<Rightarrow> nat) \<Rightarrow> bool" (infix \<open>\<prec>\<close> 50)
where
"f \<prec> g \<longleftrightarrow> f \<lesssim> g \<and> \<not> g \<lesssim> f"
--- 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 \<open>The Hebrew Alef-Bet (א-ב).\<close>
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 (\<open>א\<close>)
+ | Bet (\<open>ב\<close>)
+ | Gimel (\<open>ג\<close>)
+ | Dalet (\<open>ד\<close>)
+ | He (\<open>ה\<close>)
+ | Vav (\<open>ו\<close>)
+ | Zayin (\<open>ז\<close>)
+ | Het (\<open>ח\<close>)
+ | Tet (\<open>ט\<close>)
+ | Yod (\<open>י\<close>)
+ | Kaf (\<open>כ\<close>)
+ | Lamed (\<open>ל\<close>)
+ | Mem (\<open>מ\<close>)
+ | Nun (\<open>נ\<close>)
+ | Samekh (\<open>ס\<close>)
+ | Ayin (\<open>ע\<close>)
+ | Pe (\<open>פ\<close>)
+ | Tsadi (\<open>צ\<close>)
+ | Qof (\<open>ק\<close>)
+ | Resh (\<open>ר\<close>)
+ | Shin (\<open>ש\<close>)
+ | Tav (\<open>ת\<close>)
thm alef_bet.induct
--- 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 \<open>Definitions\<close>
locale dpo =
- fixes le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>" 50)
+ fixes le :: "['a, 'a] => bool" (infixl \<open>\<sqsubseteq>\<close> 50)
assumes refl [intro, simp]: "x \<sqsubseteq> x"
and antisym [intro]: "[| x \<sqsubseteq> y; y \<sqsubseteq> x |] ==> x = y"
and trans [trans]: "[| x \<sqsubseteq> y; y \<sqsubseteq> z |] ==> x \<sqsubseteq> z"
@@ -39,7 +39,7 @@
by (blast intro: trans)
definition
- less :: "['a, 'a] => bool" (infixl "\<sqsubset>" 50)
+ less :: "['a, 'a] => bool" (infixl \<open>\<sqsubset>\<close> 50)
where "(x \<sqsubset> y) = (x \<sqsubseteq> y & x ~= y)"
theorem abs_test:
@@ -63,11 +63,11 @@
begin
definition
- meet :: "['a, 'a] => 'a" (infixl "\<sqinter>" 70)
+ meet :: "['a, 'a] => 'a" (infixl \<open>\<sqinter>\<close> 70)
where "x \<sqinter> y = (THE inf. is_inf x y inf)"
definition
- join :: "['a, 'a] => 'a" (infixl "\<squnion>" 65)
+ join :: "['a, 'a] => 'a" (infixl \<open>\<squnion>\<close> 65)
where "x \<squnion> y = (THE sup. is_sup x y sup)"
lemma is_infI [intro?]: "i \<sqsubseteq> x \<Longrightarrow> i \<sqsubseteq> y \<Longrightarrow>
@@ -620,7 +620,7 @@
subsubsection \<open>Locale declarations and lemmas\<close>
locale Dsemi =
- fixes prod (infixl "**" 65)
+ fixes prod (infixl \<open>**\<close> 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 \<open>**\<close> 65) and one and sum (infixl \<open>+++\<close> 60) and zero +
fixes hom
assumes hom_mult [simp]: "hom (x ** y) = hom x +++ hom y"
--- 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 @@
\<close>
syntax
- "_quote" :: "'b \<Rightarrow> ('a \<Rightarrow> 'b)" ("\<guillemotleft>_\<guillemotright>" [0] 1000)
- "_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b" ("\<acute>_" [1000] 1000)
+ "_quote" :: "'b \<Rightarrow> ('a \<Rightarrow> 'b)" (\<open>\<guillemotleft>_\<guillemotright>\<close> [0] 1000)
+ "_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b" (\<open>\<acute>_\<close> [1000] 1000)
parse_translation \<open>
let
--- 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 @@
\<close>
class partial_equiv =
- fixes eqv :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sim>" 50)
+ fixes eqv :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl \<open>\<sim>\<close> 50)
assumes partial_equiv_sym [elim?]: "x \<sim> y \<Longrightarrow> y \<sim> x"
assumes partial_equiv_trans [trans]: "x \<sim> y \<Longrightarrow> y \<sim> z \<Longrightarrow> x \<sim> z"
@@ -168,7 +168,7 @@
representation of elements of a quotient type.
\<close>
-definition eqv_class :: "('a::partial_equiv) \<Rightarrow> 'a quot" ("\<lfloor>_\<rfloor>")
+definition eqv_class :: "('a::partial_equiv) \<Rightarrow> 'a quot" (\<open>\<lfloor>_\<rfloor>\<close>)
where "\<lfloor>a\<rfloor> = Abs_quot {x. a \<sim> x}"
theorem quot_rep: "\<exists>a. A = \<lfloor>a\<rfloor>"
--- 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 \<open>Serbian cyrillic letters.\<close>
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 (\<open>А\<close>)
+| azbB (\<open>Б\<close>)
+| azbV (\<open>В\<close>)
+| azbG (\<open>Г\<close>)
+| azbD (\<open>Д\<close>)
+| azbDj (\<open>Ђ\<close>)
+| azbE (\<open>Е\<close>)
+| azbZv (\<open>Ж\<close>)
+| azbZ (\<open>З\<close>)
+| azbI (\<open>И\<close>)
+| azbJ (\<open>Ј\<close>)
+| azbK (\<open>К\<close>)
+| azbL (\<open>Л\<close>)
+| azbLj (\<open>Љ\<close>)
+| azbM (\<open>М\<close>)
+| azbN (\<open>Н\<close>)
+| azbNj (\<open>Њ\<close>)
+| azbO (\<open>О\<close>)
+| azbP (\<open>П\<close>)
+| azbR (\<open>Р\<close>)
+| azbS (\<open>С\<close>)
+| azbT (\<open>Т\<close>)
+| azbC' (\<open>Ћ\<close>)
+| azbU (\<open>У\<close>)
+| azbF (\<open>Ф\<close>)
+| azbH (\<open>Х\<close>)
+| azbC (\<open>Ц\<close>)
+| azbCv (\<open>Ч\<close>)
+| azbDzv (\<open>Џ\<close>)
+| azbSv (\<open>Ш\<close>)
| azbSpc
thm azbuka.induct
text \<open>Serbian latin letters.\<close>
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 (\<open>A\<close>)
+| abcB (\<open>B\<close>)
+| abcC (\<open>C\<close>)
+| abcCv (\<open>Č\<close>)
+| abcC' (\<open>Ć\<close>)
+| abcD (\<open>D\<close>)
+| abcE (\<open>E\<close>)
+| abcF (\<open>F\<close>)
+| abcG (\<open>G\<close>)
+| abcH (\<open>H\<close>)
+| abcI (\<open>I\<close>)
+| abcJ (\<open>J\<close>)
+| abcK (\<open>K\<close>)
+| abcL (\<open>L\<close>)
+| abcM (\<open>M\<close>)
+| abcN (\<open>N\<close>)
+| abcO (\<open>O\<close>)
+| abcP (\<open>P\<close>)
+| abcR (\<open>R\<close>)
+| abcS (\<open>S\<close>)
+| abcSv (\<open>Š\<close>)
+| abcT (\<open>T\<close>)
+| abcU (\<open>U\<close>)
+| abcV (\<open>V\<close>)
+| abcZ (\<open>Z\<close>)
+| abcvZ (\<open>Ž\<close>)
| abcSpc
thm abeceda.induct
--- 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 @@
\<open>nitpick\<close>, which is used below.)
\<close>
-no_notation Groups.one_class.one ("1")
+no_notation Groups.one_class.one (\<open>1\<close>)
-datatype digit = A ("1") | B ("2") | C ("3") | D ("4") | E ("5") | F ("6") | G ("7") | H ("8") | I ("9")
+datatype digit = A (\<open>1\<close>) | B (\<open>2\<close>) | C (\<open>3\<close>) | D (\<open>4\<close>) | E (\<open>5\<close>) | F (\<open>6\<close>) | G (\<open>7\<close>) | H (\<open>8\<close>) | I (\<open>9\<close>)
definition valid :: "digit => digit => digit => digit => digit => digit => digit => digit => digit => bool" where
--- 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 \<subseteq> pset cl \<and> \<lparr>pset = S, order = induced S (order cl)\<rparr> \<in> CompleteLattice})"
-abbreviation sublat :: "['a set, 'a potype] \<Rightarrow> bool" ("_ <<= _" [51, 50] 50)
+abbreviation sublat :: "['a set, 'a potype] \<Rightarrow> bool" (\<open>_ <<= _\<close> [51, 50] 50)
where "S <<= cl \<equiv> S \<in> sublattice `` {cl}"
definition dual :: "'a potype \<Rightarrow> 'a potype"
--- 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 "\<cdot>" 60)
+ | Comb "'a trm" "'a trm" (infix \<open>\<cdot>\<close> 60)
primrec vars_of :: "'a trm \<Rightarrow> 'a set"
where
@@ -43,7 +43,7 @@
| "vars_of (Const c) = {}"
| "vars_of (M \<cdot> N) = vars_of M \<union> vars_of N"
-fun occs :: "'a trm \<Rightarrow> 'a trm \<Rightarrow> bool" (infixl "\<prec>" 54)
+fun occs :: "'a trm \<Rightarrow> 'a trm \<Rightarrow> bool" (infixl \<open>\<prec>\<close> 54)
where
"u \<prec> Var v \<longleftrightarrow> False"
| "u \<prec> Const c \<longleftrightarrow> 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 \<Rightarrow> 'a subst \<Rightarrow> 'a trm" (infixl "\<lhd>" 55)
+primrec subst :: "'a trm \<Rightarrow> 'a subst \<Rightarrow> 'a trm" (infixl \<open>\<lhd>\<close> 55)
where
"(Var v) \<lhd> s = assoc v (Var v) s"
| "(Const c) \<lhd> s = (Const c)"
| "(M \<cdot> N) \<lhd> s = (M \<lhd> s) \<cdot> (N \<lhd> s)"
-definition subst_eq (infixr "\<doteq>" 52)
+definition subst_eq (infixr \<open>\<doteq>\<close> 52)
where
"s1 \<doteq> s2 \<longleftrightarrow> (\<forall>t. t \<lhd> s1 = t \<lhd> s2)"
-fun comp :: "'a subst \<Rightarrow> 'a subst \<Rightarrow> 'a subst" (infixl "\<lozenge>" 56)
+fun comp :: "'a subst \<Rightarrow> 'a subst \<Rightarrow> 'a subst" (infixl \<open>\<lozenge>\<close> 56)
where
"[] \<lozenge> bl = bl"
| "((a,b) # al) \<lozenge> bl = (a, b \<lhd> bl) # (al \<lozenge> bl)"
--- 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 \<open>*\<close> 6)
+typedecl ('a,'b) sum (infixl \<open>+\<close> 5)
instance "fun" :: (cpo, cpo) cpo ..
instance prod :: (cpo, cpo) cpo ..
@@ -38,10 +38,10 @@
INR :: "'b \<Rightarrow> 'a+'b"
WHEN :: "['a\<Rightarrow>'c, 'b\<Rightarrow>'c, 'a+'b] \<Rightarrow> 'c"
adm :: "('a \<Rightarrow> o) \<Rightarrow> o"
- VOID :: "void" ("'(')")
- PAIR :: "['a,'b] \<Rightarrow> 'a*'b" ("(1<_,/_>)" [0,0] 100)
- COND :: "[tr,'a,'a] \<Rightarrow> 'a" ("(_ \<Rightarrow>/ (_ |/ _))" [60,60,60] 60)
- less :: "['a,'a] \<Rightarrow> o" (infixl "<<" 50)
+ VOID :: "void" (\<open>'(')\<close>)
+ PAIR :: "['a,'b] \<Rightarrow> 'a*'b" (\<open>(1<_,/_>)\<close> [0,0] 100)
+ COND :: "[tr,'a,'a] \<Rightarrow> 'a" (\<open>(_ \<Rightarrow>/ (_ |/ _))\<close> [60,60,60] 60)
+ less :: "['a,'a] \<Rightarrow> o" (infixl \<open><<\<close> 50)
axiomatization where
(** DOMAIN THEORY **)
--- 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 \<Rightarrow> prop" ("_" 5)
+judgment Trueprop :: "o \<Rightarrow> prop" (\<open>_\<close> 5)
subsection \<open>Propositional logic\<close>
-axiomatization false :: o ("\<bottom>")
+axiomatization false :: o (\<open>\<bottom>\<close>)
where falseE [elim]: "\<bottom> \<Longrightarrow> A"
-axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longrightarrow>" 25)
+axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<longrightarrow>\<close> 25)
where impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B"
and mp [dest]: "A \<longrightarrow> B \<Longrightarrow> A \<Longrightarrow> B"
-axiomatization conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35)
+axiomatization conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<and>\<close> 35)
where conjI [intro]: "A \<Longrightarrow> B \<Longrightarrow> A \<and> B"
and conjD1: "A \<and> B \<Longrightarrow> A"
and conjD2: "A \<and> B \<Longrightarrow> B"
@@ -48,20 +48,20 @@
qed
-axiomatization disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30)
+axiomatization disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<or>\<close> 30)
where disjE [elim]: "A \<or> B \<Longrightarrow> (A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C"
and disjI1 [intro]: "A \<Longrightarrow> A \<or> B"
and disjI2 [intro]: "B \<Longrightarrow> A \<or> B"
-definition true :: o ("\<top>")
+definition true :: o (\<open>\<top>\<close>)
where "\<top> \<equiv> \<bottom> \<longrightarrow> \<bottom>"
theorem trueI [intro]: \<top>
unfolding true_def ..
-definition not :: "o \<Rightarrow> o" ("\<not> _" [40] 40)
+definition not :: "o \<Rightarrow> o" (\<open>\<not> _\<close> [40] 40)
where "\<not> A \<equiv> A \<longrightarrow> \<bottom>"
theorem notI [intro]: "(A \<Longrightarrow> \<bottom>) \<Longrightarrow> \<not> A"
@@ -76,7 +76,7 @@
qed
-definition iff :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longleftrightarrow>" 25)
+definition iff :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<longleftrightarrow>\<close> 25)
where "A \<longleftrightarrow> B \<equiv> (A \<longrightarrow> B) \<and> (B \<longrightarrow> A)"
theorem iffI [intro]:
@@ -112,7 +112,7 @@
subsection \<open>Equality\<close>
-axiomatization equal :: "i \<Rightarrow> i \<Rightarrow> o" (infixl "=" 50)
+axiomatization equal :: "i \<Rightarrow> i \<Rightarrow> o" (infixl \<open>=\<close> 50)
where refl [intro]: "x = x"
and subst: "x = y \<Longrightarrow> P x \<Longrightarrow> P y"
@@ -129,11 +129,11 @@
subsection \<open>Quantifiers\<close>
-axiomatization All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10)
+axiomatization All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<forall>\<close> 10)
where allI [intro]: "(\<And>x. P x) \<Longrightarrow> \<forall>x. P x"
and allD [dest]: "\<forall>x. P x \<Longrightarrow> P a"
-axiomatization Ex :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10)
+axiomatization Ex :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>\<close> 10)
where exI [intro]: "P a \<Longrightarrow> \<exists>x. P x"
and exE [elim]: "\<exists>x. P x \<Longrightarrow> (\<And>x. P x \<Longrightarrow> C) \<Longrightarrow> C"
--- 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 \<Rightarrow> prop" ("_" 5)
+judgment Trueprop :: "o \<Rightarrow> prop" (\<open>_\<close> 5)
section \<open>Minimal logic (axiomatization)\<close>
-axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longrightarrow>" 25)
+axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<longrightarrow>\<close> 25)
where impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B"
and impE [dest, trans]: "A \<longrightarrow> B \<Longrightarrow> A \<Longrightarrow> B"
-axiomatization All :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10)
+axiomatization All :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<forall>\<close> 10)
where allI [intro]: "(\<And>x. P x) \<Longrightarrow> \<forall>x. P x"
and allE [dest]: "\<forall>x. P x \<Longrightarrow> P a"
@@ -68,7 +68,7 @@
unfolding True_def ..
-definition not :: "o \<Rightarrow> o" ("\<not> _" [40] 40)
+definition not :: "o \<Rightarrow> o" (\<open>\<not> _\<close> [40] 40)
where "not \<equiv> \<lambda>A. A \<longrightarrow> False"
lemma notI [intro]:
@@ -91,7 +91,7 @@
lemmas contradiction = notE notE' \<comment> \<open>proof by contradiction in any order\<close>
-definition conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35)
+definition conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<and>\<close> 35)
where "A \<and> B \<equiv> \<forall>C. (A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> C"
lemma conjI [intro]:
@@ -137,7 +137,7 @@
qed
-definition disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30)
+definition disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<or>\<close> 30)
where "A \<or> B \<equiv> \<forall>C. (A \<longrightarrow> C) \<longrightarrow> (B \<longrightarrow> C) \<longrightarrow> C"
lemma disjI1 [intro]:
@@ -190,7 +190,7 @@
qed
-definition Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10)
+definition Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>\<close> 10)
where "\<exists>x. P x \<equiv> \<forall>C. (\<forall>x. P x \<longrightarrow> C) \<longrightarrow> C"
lemma exI [intro]: "P a \<Longrightarrow> \<exists>x. P x"
@@ -227,14 +227,14 @@
subsubsection \<open>Extensional equality\<close>
-axiomatization equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl "=" 50)
+axiomatization equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl \<open>=\<close> 50)
where refl [intro]: "x = x"
and subst: "x = y \<Longrightarrow> P x \<Longrightarrow> P y"
-abbreviation not_equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl "\<noteq>" 50)
+abbreviation not_equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl \<open>\<noteq>\<close> 50)
where "x \<noteq> y \<equiv> \<not> (x = y)"
-abbreviation iff :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<longleftrightarrow>" 25)
+abbreviation iff :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<longleftrightarrow>\<close> 25)
where "A \<longleftrightarrow> B \<equiv> A = B"
axiomatization
@@ -401,7 +401,7 @@
axiomatization Eps :: "('a \<Rightarrow> o) \<Rightarrow> 'a"
where someI: "P x \<Longrightarrow> P (Eps P)"
-syntax "_Eps" :: "pttrn \<Rightarrow> o \<Rightarrow> 'a" ("(3SOME _./ _)" [0, 10] 10)
+syntax "_Eps" :: "pttrn \<Rightarrow> o \<Rightarrow> 'a" (\<open>(3SOME _./ _)\<close> [0, 10] 10)
syntax_consts "_Eps" \<rightleftharpoons> Eps
translations "SOME x. P" \<rightleftharpoons> "CONST Eps (\<lambda>x. P)"
--- 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] \<Rightarrow> o" (infixr "><" 35)
- limp :: "[o, o] \<Rightarrow> o" (infixr "-o" 45)
- FShriek :: "o \<Rightarrow> o" ("! _" [100] 1000)
- lconj :: "[o, o] \<Rightarrow> o" (infixr "&&" 35)
- ldisj :: "[o, o] \<Rightarrow> o" (infixr "++" 35)
- zero :: "o" ("0")
- top :: "o" ("1")
- eye :: "o" ("I")
+ tens :: "[o, o] \<Rightarrow> o" (infixr \<open>><\<close> 35)
+ limp :: "[o, o] \<Rightarrow> o" (infixr \<open>-o\<close> 45)
+ FShriek :: "o \<Rightarrow> o" (\<open>! _\<close> [100] 1000)
+ lconj :: "[o, o] \<Rightarrow> o" (infixr \<open>&&\<close> 35)
+ ldisj :: "[o, o] \<Rightarrow> o" (infixr \<open>++\<close> 35)
+ zero :: "o" (\<open>0\<close>)
+ top :: "o" (\<open>1\<close>)
+ eye :: "o" (\<open>I\<close>)
(* context manipulation *)
@@ -29,9 +29,9 @@
PromAux :: "three_seqi"
syntax
- "_Trueprop" :: "single_seqe" ("((_)/ \<turnstile> (_))" [6,6] 5)
- "_Context" :: "two_seqe" ("((_)/ :=: (_))" [6,6] 5)
- "_PromAux" :: "three_seqe" ("promaux {_||_||_}")
+ "_Trueprop" :: "single_seqe" (\<open>((_)/ \<turnstile> (_))\<close> [6,6] 5)
+ "_Context" :: "two_seqe" (\<open>((_)/ :=: (_))\<close> [6,6] 5)
+ "_PromAux" :: "three_seqe" (\<open>promaux {_||_||_}\<close>)
parse_translation \<open>
[(\<^syntax_const>\<open>_Trueprop\<close>, K (single_tr \<^const_syntax>\<open>Trueprop\<close>)),
@@ -45,10 +45,10 @@
(\<^const_syntax>\<open>PromAux\<close>, K (three_seq_tr' \<^syntax_const>\<open>_PromAux\<close>))]
\<close>
-definition liff :: "[o, o] \<Rightarrow> o" (infixr "o-o" 45)
+definition liff :: "[o, o] \<Rightarrow> o" (infixr \<open>o-o\<close> 45)
where "P o-o Q == (P -o Q) >< (Q -o P)"
-definition aneg :: "o\<Rightarrow>o" ("~_")
+definition aneg :: "o\<Rightarrow>o" (\<open>~_\<close>)
where "~A == A -o 0"
--- 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] \<Rightarrow> plf" (infixr "&" 35)
- disj :: "[plf,plf] \<Rightarrow> plf" (infixr "|" 35)
- impl :: "[plf,plf] \<Rightarrow> plf" (infixr ">" 35)
- eq :: "[plf,plf] \<Rightarrow> plf" (infixr "=" 35)
- ff :: "plf" ("ff")
+ conj :: "[plf,plf] \<Rightarrow> plf" (infixr \<open>&\<close> 35)
+ disj :: "[plf,plf] \<Rightarrow> plf" (infixr \<open>|\<close> 35)
+ impl :: "[plf,plf] \<Rightarrow> plf" (infixr \<open>>\<close> 35)
+ eq :: "[plf,plf] \<Rightarrow> plf" (infixr \<open>=\<close> 35)
+ ff :: "plf" (\<open>ff\<close>)
- PL :: "plf \<Rightarrow> o" ("[* / _ / *]" [] 100)
+ PL :: "plf \<Rightarrow> o" (\<open>[* / _ / *]\<close> [] 100)
syntax
- "_NG" :: "plf \<Rightarrow> plf" ("- _ " [50] 55)
+ "_NG" :: "plf \<Rightarrow> plf" (\<open>- _ \<close> [50] 55)
translations
"[* A & B *]" \<rightleftharpoons> "[*A*] && [*B*]"
--- 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 (\<open>0\<close>) and
Suc :: "nat \<Rightarrow> nat" and
rec :: "[nat, 'a, [nat,'a] \<Rightarrow> 'a] \<Rightarrow> 'a"
where
@@ -25,7 +25,7 @@
rec_0: "\<turnstile> rec(0,a,f) = a" and
rec_Suc: "\<turnstile> rec(Suc(m), a, f) = f(m, rec(m,a,f))"
-definition add :: "[nat, nat] \<Rightarrow> nat" (infixl "+" 60)
+definition add :: "[nat, nat] \<Rightarrow> nat" (infixl \<open>+\<close> 60)
where "m + n == rec(m, n, \<lambda>x y. Suc(y))"
--- 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] \<Rightarrow> o" (infixl "=" 50)
- Not :: "o \<Rightarrow> o" ("\<not> _" [40] 40)
- conj :: "[o,o] \<Rightarrow> o" (infixr "\<and>" 35)
- disj :: "[o,o] \<Rightarrow> o" (infixr "\<or>" 30)
- imp :: "[o,o] \<Rightarrow> o" (infixr "\<longrightarrow>" 25)
- iff :: "[o,o] \<Rightarrow> o" (infixr "\<longleftrightarrow>" 25)
- The :: "('a \<Rightarrow> o) \<Rightarrow> 'a" (binder "THE " 10)
- All :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder "\<forall>" 10)
- Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10)
+ equal :: "['a,'a] \<Rightarrow> o" (infixl \<open>=\<close> 50)
+ Not :: "o \<Rightarrow> o" (\<open>\<not> _\<close> [40] 40)
+ conj :: "[o,o] \<Rightarrow> o" (infixr \<open>\<and>\<close> 35)
+ disj :: "[o,o] \<Rightarrow> o" (infixr \<open>\<or>\<close> 30)
+ imp :: "[o,o] \<Rightarrow> o" (infixr \<open>\<longrightarrow>\<close> 25)
+ iff :: "[o,o] \<Rightarrow> o" (infixr \<open>\<longleftrightarrow>\<close> 25)
+ The :: "('a \<Rightarrow> o) \<Rightarrow> 'a" (binder \<open>THE \<close> 10)
+ All :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<forall>\<close> 10)
+ Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>\<close> 10)
syntax
- "_Trueprop" :: "two_seqe" ("((_)/ \<turnstile> (_))" [6,6] 5)
+ "_Trueprop" :: "two_seqe" (\<open>((_)/ \<turnstile> (_))\<close> [6,6] 5)
parse_translation \<open>[(\<^syntax_const>\<open>_Trueprop\<close>, K (two_seq_tr \<^const_syntax>\<open>Trueprop\<close>))]\<close>
print_translation \<open>[(\<^const_syntax>\<open>Trueprop\<close>, K (two_seq_tr' \<^syntax_const>\<open>_Trueprop\<close>))]\<close>
abbreviation
- not_equal (infixl "\<noteq>" 50) where
+ not_equal (infixl \<open>\<noteq>\<close> 50) where
"x \<noteq> y \<equiv> \<not> (x = y)"
axiomatization where
@@ -104,7 +104,7 @@
The: "\<lbrakk>$H \<turnstile> $E, P(a), $F; \<And>x.$H, P(x) \<turnstile> $E, x=a, $F\<rbrakk> \<Longrightarrow>
$H \<turnstile> $E, P(THE x. P(x)), $F"
-definition If :: "[o, 'a, 'a] \<Rightarrow> 'a" ("(if (_)/ then (_)/ else (_))" 10)
+definition If :: "[o, 'a, 'a] \<Rightarrow> 'a" (\<open>(if (_)/ then (_)/ else (_))\<close> 10)
where "If(P,x,y) \<equiv> THE z::'a. (P \<longrightarrow> z = x) \<and> (\<not> P \<longrightarrow> z = y)"
--- 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 \<open>modal.ML\<close>
consts
- box :: "o\<Rightarrow>o" ("[]_" [50] 50)
- dia :: "o\<Rightarrow>o" ("<>_" [50] 50)
+ box :: "o\<Rightarrow>o" (\<open>[]_\<close> [50] 50)
+ dia :: "o\<Rightarrow>o" (\<open><>_\<close> [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" (\<open>(_)|L>(_)\<close> [6,6] 5)
+ "_Rstar" :: "two_seqe" (\<open>(_)|R>(_)\<close> [6,6] 5)
ML \<open>
fun star_tr c [s1, s2] = Syntax.const c $ seq_tr s1 $ seq_tr s2;
@@ -34,10 +34,10 @@
(\<^const_syntax>\<open>Rstar\<close>, K (star_tr' \<^syntax_const>\<open>_Rstar\<close>))]
\<close>
-definition strimp :: "[o,o]\<Rightarrow>o" (infixr "--<" 25)
+definition strimp :: "[o,o]\<Rightarrow>o" (infixr \<open>--<\<close> 25)
where "P --< Q == [](P \<longrightarrow> Q)"
-definition streqv :: "[o,o]\<Rightarrow>o" (infixr ">-<" 25)
+definition streqv :: "[o,o]\<Rightarrow>o" (infixr \<open>>-<\<close> 25)
where "P >-< Q == (P --< Q) \<and> (Q --< P)"
--- 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'\<Rightarrow>seq', seq'\<Rightarrow>seq', seq'\<Rightarrow>seq'] \<Rightarrow> prop"
syntax
"_S43pi" :: "[seq, seq, seq, seq, seq, seq] \<Rightarrow> prop"
- ("S43pi((_);(_);(_);(_);(_);(_))" [] 5)
+ (\<open>S43pi((_);(_);(_);(_);(_);(_))\<close> [] 5)
parse_translation \<open>
let
--- 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] \<Rightarrow> seq" ("__")
+ "_SeqEmp" :: seq (\<open>\<close>)
+ "_SeqApp" :: "[seqobj,seqcont] \<Rightarrow> seq" (\<open>__\<close>)
- "_SeqContEmp" :: seqcont ("")
- "_SeqContApp" :: "[seqobj,seqcont] \<Rightarrow> seqcont" (",/ __")
+ "_SeqContEmp" :: seqcont (\<open>\<close>)
+ "_SeqContApp" :: "[seqobj,seqcont] \<Rightarrow> seqcont" (\<open>,/ __\<close>)
- "_SeqO" :: "o \<Rightarrow> seqobj" ("_")
- "_SeqId" :: "'a \<Rightarrow> seqobj" ("$_")
+ "_SeqO" :: "o \<Rightarrow> seqobj" (\<open>_\<close>)
+ "_SeqId" :: "'a \<Rightarrow> seqobj" (\<open>$_\<close>)
type_synonym single_seqe = "[seq,seqobj] \<Rightarrow> prop"
type_synonym single_seqi = "[seq'\<Rightarrow>seq',seq'\<Rightarrow>seq'] \<Rightarrow> prop"
@@ -52,7 +52,7 @@
syntax
(*Constant to allow definitions of SEQUENCES of formulas*)
- "_Side" :: "seq\<Rightarrow>(seq'\<Rightarrow>seq')" ("<<(_)>>")
+ "_Side" :: "seq\<Rightarrow>(seq'\<Rightarrow>seq')" (\<open><<(_)>>\<close>)
ML \<open>