src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy
changeset 80914 d97fdabd9e2b
parent 69597 ff784d5a5bfb
equal deleted inserted replaced
80913:46f59511b7bb 80914:d97fdabd9e2b
   208 
   208 
   209 subsection \<open>Lambda\<close>
   209 subsection \<open>Lambda\<close>
   210 
   210 
   211 datatype type =
   211 datatype type =
   212     Atom nat
   212     Atom nat
   213   | Fun type type    (infixr "\<Rightarrow>" 200)
   213   | Fun type type    (infixr \<open>\<Rightarrow>\<close> 200)
   214 
   214 
   215 datatype dB =
   215 datatype dB =
   216     Var nat
   216     Var nat
   217   | App dB dB (infixl "\<degree>" 200)
   217   | App dB dB (infixl \<open>\<degree>\<close> 200)
   218   | Abs type dB
   218   | Abs type dB
   219 
   219 
   220 primrec
   220 primrec
   221   nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
   221   nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" (\<open>_\<langle>_\<rangle>\<close> [90, 0] 91)
   222 where
   222 where
   223   "[]\<langle>i\<rangle> = None"
   223   "[]\<langle>i\<rangle> = None"
   224 | "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
   224 | "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
   225 
   225 
   226 inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
   226 inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
   227 where
   227 where
   228   "nth_el' (x # xs) 0 x"
   228   "nth_el' (x # xs) 0 x"
   229 | "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
   229 | "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
   230 
   230 
   231 inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
   231 inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  (\<open>_ \<turnstile> _ : _\<close> [50, 50, 50] 50)
   232   where
   232   where
   233     Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
   233     Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
   234   | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
   234   | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
   235   | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
   235   | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
   236 
   236 
   240     "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
   240     "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
   241   | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
   241   | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
   242   | "lift (Abs T s) k = Abs T (lift s (k + 1))"
   242   | "lift (Abs T s) k = Abs T (lift s (k + 1))"
   243 
   243 
   244 primrec
   244 primrec
   245   subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
   245   subst :: "[dB, dB, nat] => dB"  (\<open>_[_'/_]\<close> [300, 0, 0] 300)
   246 where
   246 where
   247     subst_Var: "(Var i)[s/k] =
   247     subst_Var: "(Var i)[s/k] =
   248       (if k < i then Var (i - 1) else if i = k then s else Var i)"
   248       (if k < i then Var (i - 1) else if i = k then s else Var i)"
   249   | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
   249   | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
   250   | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
   250   | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
   251 
   251 
   252 inductive beta :: "[dB, dB] => bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
   252 inductive beta :: "[dB, dB] => bool"  (infixl \<open>\<rightarrow>\<^sub>\<beta>\<close> 50)
   253   where
   253   where
   254     beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
   254     beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
   255   | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
   255   | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
   256   | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
   256   | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
   257   | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
   257   | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"