src/HOL/Predicate_Compile_Examples/Lambda_Example.thy
changeset 80914 d97fdabd9e2b
parent 69597 ff784d5a5bfb
--- 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"