Nominal/Ex/SFT/Lambda.thy
changeset 2893 589b1a0c75e6
child 2894 8ec94871de1e
equal deleted inserted replaced
2892:a9f3600c9ae6 2893:589b1a0c75e6
       
     1 header {* Definition of Lambda terms and convertibility *}
       
     2 
       
     3 theory Lambda imports Nominal2 begin
       
     4 
       
     5 lemma [simp]: "supp x = {} \<Longrightarrow> y \<sharp> x"
       
     6   unfolding fresh_def by blast
       
     7 
       
     8 atom_decl var
       
     9 
       
    10 nominal_datatype lam =
       
    11   V "var"
       
    12 | Ap "lam" "lam" (infixl "\<cdot>" 98)
       
    13 | Lm x::"var" l::"lam"  bind x in l ("\<integral> _. _" [97, 97] 99)
       
    14 
       
    15 nominal_primrec
       
    16   subst :: "lam \<Rightarrow> var \<Rightarrow> lam \<Rightarrow> lam"  ("_ [_ ::= _]" [90, 90, 90] 90)
       
    17 where
       
    18   "(V x)[y ::= s] = (if x = y then s else (V x))"
       
    19 | "(t1 \<cdot> t2)[y ::= s] = (t1[y ::= s]) \<cdot> (t2[y ::= s])"
       
    20 | "atom x \<sharp> (y, s) \<Longrightarrow> (\<integral>x. t)[y ::= s] = \<integral>x.(t[y ::= s])"
       
    21 proof auto
       
    22   fix a b :: lam and aa :: var and P
       
    23   assume "\<And>x y s. a = V x \<and> aa = y \<and> b = s \<Longrightarrow> P"
       
    24     "\<And>t1 t2 y s. a = t1 \<cdot> t2 \<and> aa = y \<and> b = s \<Longrightarrow> P"
       
    25     "\<And>x y s t. \<lbrakk>atom x \<sharp> (y, s); a = \<integral> x. t \<and> aa = y \<and> b = s\<rbrakk> \<Longrightarrow> P"
       
    26   then show "P"
       
    27     by (rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust)
       
    28        (blast, blast, simp add: fresh_star_def)
       
    29 next
       
    30   fix x :: var and t and xa :: var and ya sa ta
       
    31   assume *: "eqvt_at subst_sumC (t, ya, sa)"
       
    32     "atom x \<sharp> (ya, sa)" "atom xa \<sharp> (ya, sa)"
       
    33     "[[atom x]]lst. t = [[atom xa]]lst. ta"
       
    34   then show "[[atom x]]lst. subst_sumC (t, ya, sa) = [[atom xa]]lst. subst_sumC (ta, ya, sa)"
       
    35     apply -
       
    36     apply (erule Abs_lst1_fcb)
       
    37     apply(simp (no_asm) add: Abs_fresh_iff)
       
    38     apply(drule_tac a="atom xa" in fresh_eqvt_at)
       
    39     apply(simp add: finite_supp)
       
    40     apply(simp_all add: fresh_Pair_elim Abs_fresh_iff Abs1_eq_iff)
       
    41     apply(subgoal_tac "(atom x \<rightleftharpoons> atom xa) \<bullet> ya = ya")
       
    42     apply(subgoal_tac "(atom x \<rightleftharpoons> atom xa) \<bullet> sa = sa")
       
    43     apply(simp add: atom_eqvt eqvt_at_def)
       
    44     apply(rule perm_supp_eq, simp add: supp_swap fresh_star_def fresh_Pair)+
       
    45     done
       
    46 next
       
    47   show "eqvt subst_graph" unfolding eqvt_def subst_graph_def
       
    48     by (rule, perm_simp, rule)
       
    49 qed
       
    50 
       
    51 termination
       
    52   by (relation "measure (\<lambda>(t,_,_). size t)")
       
    53      (simp_all add: lam.size)
       
    54 
       
    55 lemma subst4[eqvt]:
       
    56   shows "(p \<bullet> t[x ::= s]) = (p \<bullet> t)[(p \<bullet> x) ::= (p \<bullet> s)]"
       
    57   by (induct t x s rule: subst.induct) (simp_all)
       
    58 
       
    59 lemma subst1[simp]:
       
    60   shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t"
       
    61   by (nominal_induct t avoiding: x s rule: lam.strong_induct)
       
    62      (auto simp add: lam.fresh fresh_at_base)
       
    63 
       
    64 lemma subst2[simp]: "supp t = {} \<Longrightarrow> t[x ::= s] = t"
       
    65   by (simp add: fresh_def)
       
    66 
       
    67 lemma subst3[simp]: "M [x ::= V x] = M"
       
    68   by (rule_tac lam="M" and c="x" in lam.strong_induct)
       
    69      (simp_all add: fresh_star_def lam.fresh fresh_Pair)
       
    70 
       
    71 inductive
       
    72   beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infix "\<approx>" 80)
       
    73 where
       
    74   bI: "(\<integral>x. M) \<cdot> N \<approx> M[x ::= N]"
       
    75 | b1: "M \<approx> M"
       
    76 | b2: "M \<approx> N \<Longrightarrow> N \<approx> M"
       
    77 | b3: "M \<approx> N \<Longrightarrow> N \<approx> L \<Longrightarrow> M \<approx> L"
       
    78 | b4: "M \<approx> N \<Longrightarrow> Z \<cdot> M \<approx> Z \<cdot> N"
       
    79 | b5: "M \<approx> N \<Longrightarrow> M \<cdot> Z \<approx> N \<cdot> Z"
       
    80 | b6: "M \<approx> N \<Longrightarrow> \<integral>x. M \<approx> \<integral>x. N"
       
    81 
       
    82 lemmas [trans] = b3
       
    83 equivariance beta
       
    84 
       
    85 end