diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Predicate_Compile_Examples/Examples.thy --- a/src/HOL/Predicate_Compile_Examples/Examples.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Predicate_Compile_Examples/Examples.thy Fri Sep 20 19:51:08 2024 +0200 @@ -166,15 +166,15 @@ datatype type = Atom nat - | Fun type type (infixr "\" 200) + | Fun type type (infixr \\\ 200) datatype dB = Var nat - | App dB dB (infixl "\" 200) + | App dB dB (infixl \\\ 200) | Abs type dB primrec - nth_el :: "'a list \ nat \ 'a option" ("_\_\" [90, 0] 91) + nth_el :: "'a list \ nat \ 'a option" (\_\_\\ [90, 0] 91) where "[]\i\ = None" | "(x # xs)\i\ = (case i of 0 \ Some x | Suc j \ xs \j\)" @@ -184,7 +184,7 @@ "nth_el' (x # xs) 0 x" | "nth_el' xs i y \ nth_el' (x # xs) (Suc i) y" -inductive typing :: "type list \ dB \ type \ bool" ("_ \ _ : _" [50, 50, 50] 50) +inductive typing :: "type list \ dB \ type \ bool" (\_ \ _ : _\ [50, 50, 50] 50) where Var [intro!]: "nth_el' env x T \ env \ Var x : T" | Abs [intro!]: "T # env \ t : U \ env \ Abs T t : (T \ U)" @@ -198,14 +198,14 @@ | "lift (Abs T s) k = Abs T (lift s (k + 1))" primrec - subst :: "[dB, dB, nat] => dB" ("_[_'/_]" [300, 0, 0] 300) + subst :: "[dB, dB, nat] => dB" (\_[_'/_]\ [300, 0, 0] 300) where subst_Var: "(Var i)[s/k] = (if k < i then Var (i - 1) else if i = k then s else Var i)" | subst_App: "(t \ u)[s/k] = t[s/k] \ u[s/k]" | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])" -inductive beta :: "[dB, dB] => bool" (infixl "\\<^sub>\" 50) +inductive beta :: "[dB, dB] => bool" (infixl \\\<^sub>\\ 50) where beta [simp, intro!]: "Abs T s \ t \\<^sub>\ s[t/0]" | appL [simp, intro!]: "s \\<^sub>\ t ==> s \ u \\<^sub>\ t \ u" @@ -284,13 +284,13 @@ datatype com' = SKIP - | Assign name aexp ("_ ::= _" [1000, 61] 61) - | Semi com' com' ("_; _" [60, 61] 60) - | If bexp com' com' ("IF _ THEN _ ELSE _" [0, 0, 61] 61) - | While bexp com' ("WHILE _ DO _" [0, 61] 61) + | Assign name aexp (\_ ::= _\ [1000, 61] 61) + | Semi com' com' (\_; _\ [60, 61] 60) + | If bexp com' com' (\IF _ THEN _ ELSE _\ [0, 0, 61] 61) + | While bexp com' (\WHILE _ DO _\ [0, 61] 61) inductive - big_step :: "com' * state' \ state' \ bool" (infix "\" 55) + big_step :: "com' * state' \ state' \ bool" (infix \\\ 55) where Skip: "(SKIP,s) \ s" | Assign: "(x ::= a,s) \ s(x := aval a s)"