# HG changeset patch # User Christian Urban # Date 1299575269 0 # Node ID 651355113eeec3339315396f7667ee19c889f4e2 # Parent a9e63abf3febc9f8eda8d120111a36a997a5c285# Parent b5ffcb30b6d0252ff7548688319f1f5998b4140d merged diff -r a9e63abf3feb -r 651355113eee Nominal/Ex/Lambda_add.thy --- a/Nominal/Ex/Lambda_add.thy Tue Mar 08 09:07:27 2011 +0000 +++ b/Nominal/Ex/Lambda_add.thy Tue Mar 08 09:07:49 2011 +0000 @@ -73,6 +73,80 @@ apply (simp add: lam.eq_iff Abs1_eq_iff lam.supp fresh_def supp_at_base) done +nominal_primrec + Pair_pre +where + "atom z \ (M, N) \ (Pair_pre M N) (Lam [z]. _) = (Lam [z]. App (App (Var z) M) N)" +| "(Pair_pre M N) (App l r) = (App l r)" +| "(Pair_pre M N) (Var v) = (Var v)" +apply (subgoal_tac "\p a x. Pair_pre_graph a x \ Pair_pre_graph (p \ a) (p \ x)") +apply (simp add: eqvt_def) +apply (intro allI) +apply(simp add: permute_fun_def) +apply(rule ext) +apply(rule ext) +apply(simp add: permute_bool_def) +apply rule +apply(drule_tac x="p" in meta_spec) +apply(drule_tac x="- p \ x" in meta_spec) +apply(drule_tac x="- p \ xa" in meta_spec) +apply simp +apply(drule_tac x="-p" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply simp +apply (erule Pair_pre_graph.induct) +apply (simp_all add: Pair_pre_graph.intros)[3] + apply (case_tac x) + apply clarify + apply (simp add: symmetric[OF atomize_conjL]) + apply (rule_tac y="c" and c="(ab, ba)" in lam.strong_exhaust) + apply (simp_all add: lam.distinct symmetric[OF atomize_conjL] lam.eq_iff Abs1_eq_iff fresh_star_def) + apply auto[1] + apply (simp_all add: fresh_Pair lam.fresh fresh_at_base) + apply (rule swap_fresh_fresh[symmetric]) + prefer 3 + apply (rule swap_fresh_fresh[symmetric]) + apply simp_all + done + +termination + by (relation "measure (\(t,_). size t)") + (simp_all add: lam.size) + +consts cx :: name + +fun + Pair where + "(Pair M N) = (Pair_pre M N (Lam [cx]. Var cx))" + +lemma Pair : "atom z \ (M, N) \ Pair M N = (Lam [z]. App (App (Var z) M) N)" + apply simp + apply (subgoal_tac "Lam [cx]. Var cx = Lam[z]. Var z") + apply simp + apply (auto simp add: lam.eq_iff Abs1_eq_iff) + by (metis Rep_name_inject atom_name_def fresh_at_base lam.fresh(1)) + + +definition Lam_I_pre : "\ \ Lam[cx]. Var cx" + +lemma Lam_I : "\ = Lam[x]. Var x" + by (simp add: Lam_I_pre lam.eq_iff Abs1_eq_iff lam.fresh supp_at_base) + (metis Rep_name_inverse atom_name_def fresh_at_base) + +definition c1 :: name where "c1 \ Abs_name (Atom (Sort ''Lambda.name'' []) 1)" +definition c2 :: name where "c2 \ Abs_name (Atom (Sort ''Lambda.name'' []) 2)" + +lemma c1_c2[simp]: + "c1 \ c2" + unfolding c1_def c2_def + by (simp add: Abs_name_inject) + +definition Lam_F_pre : "\ \ Lam[c1]. Lam[c2]. Var c2" +lemma Lam_F : "\ = Lam[x]. Lam[y]. Var y" + by (simp add: Lam_F_pre lam.eq_iff Abs1_eq_iff lam.fresh supp_at_base) + (metis Rep_name_inverse atom_name_def fresh_at_base) + end