diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/Height.thy --- a/src/HOL/Nominal/Examples/Height.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/Height.thy Fri Sep 20 19:51:08 2024 +0200 @@ -12,7 +12,7 @@ nominal_datatype lam = Var "name" | App "lam" "lam" - | Lam "\name\lam" ("Lam [_]._" [100,100] 100) + | Lam "\name\lam" (\Lam [_]._\ [100,100] 100) text \Definition of the height-function on lambda-terms.\ @@ -31,7 +31,7 @@ text \Definition of capture-avoiding substitution.\ nominal_primrec - subst :: "lam \ name \ lam \ lam" ("_[_::=_]" [100,100,100] 100) + subst :: "lam \ name \ lam \ lam" (\_[_::=_]\ [100,100,100] 100) where "(Var x)[y::=t'] = (if x=y then t' else (Var x))" | "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"