diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Nominal/Examples/SN.thy --- a/src/HOL/Nominal/Examples/SN.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Nominal/Examples/SN.thy Fri Sep 20 19:51:08 2024 +0200 @@ -71,7 +71,7 @@ (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh') inductive - Beta :: "lam\lam\bool" (" _ \\<^sub>\ _" [80,80] 80) + Beta :: "lam\lam\bool" (\ _ \\<^sub>\ _\ [80,80] 80) where b1[intro!]: "s1 \\<^sub>\ s2 \ App s1 t \\<^sub>\ App s2 t" | b2[intro!]: "s1\\<^sub>\s2 \ App t s1 \\<^sub>\ App t s2" @@ -113,7 +113,7 @@ nominal_datatype ty = TVar "nat" - | TArr "ty" "ty" (infix "\" 200) + | TArr "ty" "ty" (infix \\\ 200) lemma fresh_ty: fixes a ::"name" @@ -144,7 +144,7 @@ (auto simp add: fresh_prod fresh_list_cons fresh_atm) inductive - typing :: "(name\ty) list\lam\ty\bool" ("_ \ _ : _" [60,60,60] 60) + typing :: "(name\ty) list\lam\ty\bool" (\_ \ _ : _\ [60,60,60] 60) where t1[intro]: "\valid \; (a,\)\set \\ \ \ \ Var a : \" | t2[intro]: "\\ \ t1 : \\\; \ \ t2 : \\ \ \ \ App t1 t2 : \" @@ -230,7 +230,7 @@ (* a slight hack to get the first element of applications *) (* this is needed to get (SN t) from SN (App t s) *) inductive - FST :: "lam\lam\bool" (" _ \ _" [80,80] 80) + FST :: "lam\lam\bool" (\ _ \ _\ [80,80] 80) where fst[intro!]: "(App t s) \ t" @@ -455,12 +455,12 @@ qed abbreviation - mapsto :: "(name\lam) list \ name \ lam \ bool" ("_ maps _ to _" [55,55,55] 55) + mapsto :: "(name\lam) list \ name \ lam \ bool" (\_ maps _ to _\ [55,55,55] 55) where "\ maps x to e \ (lookup \ x) = e" abbreviation - closes :: "(name\lam) list \ (name\ty) list \ bool" ("_ closes _" [55,55] 55) + closes :: "(name\lam) list \ (name\ty) list \ bool" (\_ closes _\ [55,55] 55) where "\ closes \ \ \x T. ((x,T) \ set \ \ (\t. \ maps x to t \ t \ RED T))"