Nominal/Ex/LetRec2.thy
branchNominal2-Isabelle2011-1
changeset 3071 11f6a561eb4b
parent 3029 6fd3fc3254ee
equal deleted inserted replaced
3070:4b4742aa43f2 3071:11f6a561eb4b
     3 begin
     3 begin
     4 
     4 
     5 atom_decl name
     5 atom_decl name
     6 
     6 
     7 nominal_datatype trm =
     7 nominal_datatype trm =
     8   Var "name"
     8   Vr "name"
     9 | App "trm" "trm"
     9 | Ap "trm" "trm"
    10 | Lam x::"name" t::"trm"  binds x in t
    10 | Lm x::"name" t::"trm"  binds (set) x in t
    11 | Let as::"assn" t::"trm"   binds "bn as" in t
    11 | Lt a::"lts" t::"trm"   binds "bn a" in a t
    12 | Let_rec as::"assn" t::"trm"   binds "bn as" in as t
    12 and lts =
    13 and assn =
    13   Lnil
    14   ANil
    14 | Lcons "name" "trm" "lts"
    15 | ACons "name" "trm" "assn"
       
    16 binder
    15 binder
    17   bn
    16   bn
    18 where
    17 where
    19   "bn (ANil) = []"
    18   "bn Lnil = []"
    20 | "bn (ACons x t as) = (atom x) # (bn as)"
    19 | "bn (Lcons x t l) = (atom x) # (bn l)"
    21 
       
    22 thm trm_assn.eq_iff
       
    23 thm trm_assn.supp
       
    24 
       
    25 thm trm_assn.fv_defs
       
    26 thm trm_assn.eq_iff
       
    27 thm trm_assn.bn_defs
       
    28 thm trm_assn.perm_simps
       
    29 thm trm_assn.induct
       
    30 thm trm_assn.distinct
       
    31 
    20 
    32 
    21 
    33 
    22 thm trm_lts.fv_defs
    34 section {* Tests *}
    23 thm trm_lts.eq_iff
       
    24 thm trm_lts.bn_defs
       
    25 thm trm_lts.perm_simps
       
    26 thm trm_lts.induct
       
    27 thm trm_lts.distinct
    35 
    28 
    36 
    29 
    37 (* why is this not in HOL simpset? *)
       
    38 (*
       
    39 lemma set_sub: "{a, b} - {b} = {a} - {b}"
       
    40 by auto
       
    41 
       
    42 lemma lets_bla:
       
    43   "x \<noteq> z \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> y \<Longrightarrow>(Lt (Lcons x (Vr y) Lnil) (Vr x)) \<noteq> (Lt (Lcons x (Vr z) Lnil) (Vr x))"
       
    44   apply (auto simp add: trm_lts.eq_iff alphas set_sub supp_at_base)
       
    45   done
       
    46 
       
    47 lemma lets_ok:
       
    48   "(Lt (Lcons x (Vr x) Lnil) (Vr x)) = (Lt (Lcons y (Vr y) Lnil) (Vr y))"
       
    49   apply (simp add: trm_lts.eq_iff)
       
    50   apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
       
    51   apply (simp_all add: alphas fresh_star_def eqvts supp_at_base)
       
    52   done
       
    53 
       
    54 lemma lets_ok3:
       
    55   "x \<noteq> y \<Longrightarrow>
       
    56    (Lt (Lcons x (Ap (Vr y) (Vr x)) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) \<noteq>
       
    57    (Lt (Lcons y (Ap (Vr x) (Vr y)) (Lcons x (Vr x) Lnil)) (Ap (Vr x) (Vr y)))"
       
    58   apply (simp add: alphas trm_lts.eq_iff)
       
    59   done
       
    60 
       
    61 
       
    62 lemma lets_not_ok1:
       
    63   "x \<noteq> y \<Longrightarrow>
       
    64    (Lt (Lcons x (Vr x) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) \<noteq>
       
    65    (Lt (Lcons y (Vr x) (Lcons x (Vr y) Lnil)) (Ap (Vr x) (Vr y)))"
       
    66   apply (simp add: alphas trm_lts.eq_iff)
       
    67   done
       
    68 
       
    69 lemma lets_nok:
       
    70   "x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow>
       
    71    (Lt (Lcons x (Ap (Vr z) (Vr z)) (Lcons y (Vr z) Lnil)) (Ap (Vr x) (Vr y))) \<noteq>
       
    72    (Lt (Lcons y (Vr z) (Lcons x (Ap (Vr z) (Vr z)) Lnil)) (Ap (Vr x) (Vr y)))"
       
    73   apply (simp add: alphas trm_lts.eq_iff fresh_star_def)
       
    74   done
       
    75 
       
    76 lemma lets_ok4:
       
    77   "(Lt (Lcons x (Ap (Vr y) (Vr x)) (Lcons y (Vr y) Lnil)) (Ap (Vr x) (Vr y))) =
       
    78    (Lt (Lcons y (Ap (Vr x) (Vr y)) (Lcons x (Vr x) Lnil)) (Ap (Vr y) (Vr x)))"
       
    79   apply (simp add: alphas trm_lts.eq_iff supp_at_base)
       
    80   apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
       
    81   apply (simp add: atom_eqvt fresh_star_def)
       
    82   done
       
    83 *)
       
    84 end
    30 end
    85 
    31 
    86 
    32 
    87 
    33