diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Contexts.thy --- a/src/HOL/Nominal/Examples/Contexts.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Contexts.thy Fri Sep 20 19:51:08 2024 +0200 @@ -25,7 +25,7 @@ nominal_datatype lam = Var "name" | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + | Lam "\name\lam" (\Lam [_]._\ [100,100] 100) text \ Contexts - the lambda case in contexts does not bind @@ -33,15 +33,15 @@ \ nominal_datatype ctx = - Hole ("\" 1000) + Hole (\\\ 1000) | CAppL "ctx" "lam" | CAppR "lam" "ctx" - | CLam "name" "ctx" ("CLam [_]._" [100,100] 100) + | CLam "name" "ctx" (\CLam [_]._\ [100,100] 100) text \Capture-Avoiding Substitution\ nominal_primrec - subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\ [100,100,100] 100) where "(Var x)[y::=s] = (if x=y then s else (Var x))" | "(App t\<^sub>1 t\<^sub>2)[y::=s] = App (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])" @@ -58,7 +58,7 @@ \ nominal_primrec - filling :: "ctx \ lam \ lam" ("_\_\" [100,100] 100) + filling :: "ctx \ lam \ lam" (\_\_\\ [100,100] 100) where "\\t\ = t" | "(CAppL E t')\t\ = App (E\t\) t'" @@ -79,7 +79,7 @@ text \The composition of two contexts.\ nominal_primrec - ctx_compose :: "ctx \ ctx \ ctx" ("_ \ _" [100,100] 100) + ctx_compose :: "ctx \ ctx \ ctx" (\_ \ _\ [100,100] 100) where "\ \ E' = E'" | "(CAppL E t') \ E' = CAppL (E \ E') t'" @@ -94,14 +94,14 @@ text \Beta-reduction via contexts in the Felleisen-Hieb style.\ inductive - ctx_red :: "lam\lam\bool" ("_ \x _" [80,80] 80) + ctx_red :: "lam\lam\bool" (\_ \x _\ [80,80] 80) where xbeta[intro]: "E\App (Lam [x].t) t'\ \x E\t[x::=t']\" text \Beta-reduction via congruence rules in the Plotkin style.\ inductive - cong_red :: "lam\lam\bool" ("_ \o _" [80,80] 80) + cong_red :: "lam\lam\bool" (\_ \o _\ [80,80] 80) where obeta[intro]: "App (Lam [x].t) t' \o t[x::=t']" | oapp1[intro]: "t \o t' \ App t t2 \o App t' t2"