src/HOL/Predicate_Compile_Examples/Lambda_Example.thy
changeset 80914 d97fdabd9e2b
parent 69597 ff784d5a5bfb
equal deleted inserted replaced
80913:46f59511b7bb 80914:d97fdabd9e2b
     4 
     4 
     5 subsection \<open>Lambda\<close>
     5 subsection \<open>Lambda\<close>
     6 
     6 
     7 datatype type =
     7 datatype type =
     8     Atom nat
     8     Atom nat
     9   | Fun type type    (infixr "\<Rightarrow>" 200)
     9   | Fun type type    (infixr \<open>\<Rightarrow>\<close> 200)
    10 
    10 
    11 datatype dB =
    11 datatype dB =
    12     Var nat
    12     Var nat
    13   | App dB dB (infixl "\<degree>" 200)
    13   | App dB dB (infixl \<open>\<degree>\<close> 200)
    14   | Abs type dB
    14   | Abs type dB
    15 
    15 
    16 primrec
    16 primrec
    17   nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
    17   nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" (\<open>_\<langle>_\<rangle>\<close> [90, 0] 91)
    18 where
    18 where
    19   "[]\<langle>i\<rangle> = None"
    19   "[]\<langle>i\<rangle> = None"
    20 | "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
    20 | "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
    21 
    21 
    22 inductive nth_el1 :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
    22 inductive nth_el1 :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
    23 where
    23 where
    24   "nth_el1 (x # xs) 0 x"
    24   "nth_el1 (x # xs) 0 x"
    25 | "nth_el1 xs i y \<Longrightarrow> nth_el1 (x # xs) (Suc i) y"
    25 | "nth_el1 xs i y \<Longrightarrow> nth_el1 (x # xs) (Suc i) y"
    26 
    26 
    27 inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
    27 inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  (\<open>_ \<turnstile> _ : _\<close> [50, 50, 50] 50)
    28   where
    28   where
    29     Var [intro!]: "nth_el1 env x T \<Longrightarrow> env \<turnstile> Var x : T"
    29     Var [intro!]: "nth_el1 env x T \<Longrightarrow> env \<turnstile> Var x : T"
    30   | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
    30   | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
    31   | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
    31   | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
    32 
    32 
    40 primrec pred :: "nat => nat"
    40 primrec pred :: "nat => nat"
    41 where
    41 where
    42   "pred (Suc i) = i"
    42   "pred (Suc i) = i"
    43 
    43 
    44 primrec
    44 primrec
    45   subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
    45   subst :: "[dB, dB, nat] => dB"  (\<open>_[_'/_]\<close> [300, 0, 0] 300)
    46 where
    46 where
    47     subst_Var: "(Var i)[s/k] =
    47     subst_Var: "(Var i)[s/k] =
    48       (if k < i then Var (pred i) else if i = k then s else Var i)"
    48       (if k < i then Var (pred i) else if i = k then s else Var i)"
    49   | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
    49   | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
    50   | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
    50   | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
    51 
    51 
    52 inductive beta :: "[dB, dB] => bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
    52 inductive beta :: "[dB, dB] => bool"  (infixl \<open>\<rightarrow>\<^sub>\<beta>\<close> 50)
    53   where
    53   where
    54     beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
    54     beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
    55   | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
    55   | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
    56   | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
    56   | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
    57   | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
    57   | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"