--- a/Nominal/Ex/Let.thy Fri Feb 04 03:52:38 2011 +0000
+++ b/Nominal/Ex/Let.thy Sat Feb 05 07:38:22 2011 +0900
@@ -73,11 +73,96 @@
apply (simp add: alphas trm_assn.supp supp_at_base x y fresh_star_def atom_eqvt)
by (metis Rep_name_inverse atom_name_def flip_fresh_fresh fresh_atom fresh_perm x y)
-lemma
- assumes "atom a \<noteq> atom c \<and> atom b \<noteq> atom c"
- shows "\<exists>p. ([atom a], Var c) \<approx>lst (op =) supp p ([atom b], Var c) \<and> supp p \<subseteq> (supp (Var c) \<inter> {atom a}) \<union> (supp (Var c) \<inter> {atom b})"
- apply (simp add: alphas trm_assn.supp supp_at_base assms trm_assn.eq_iff atom_eqvt fresh_star_def)
- oops
+lemma alpha_bn_refl: "alpha_bn x x"
+apply (induct x rule: trm_assn.inducts(2))
+apply (rule TrueI)
+apply (auto simp add: trm_assn.eq_iff)
+done
+
+lemma alpha_bn_inducts_raw:
+ "\<lbrakk>alpha_bn_raw a b; P3 ANil_raw ANil_raw;
+ \<And>trm_raw trm_rawa assn_raw assn_rawa name namea.
+ \<lbrakk>alpha_trm_raw trm_raw trm_rawa; alpha_bn_raw assn_raw assn_rawa;
+ P3 assn_raw assn_rawa\<rbrakk>
+ \<Longrightarrow> P3 (ACons_raw name trm_raw assn_raw)
+ (ACons_raw namea trm_rawa assn_rawa)\<rbrakk> \<Longrightarrow> P3 a b"
+ by (erule alpha_trm_raw_alpha_assn_raw_alpha_bn_raw.inducts(3)[of _ _ "\<lambda>x y. True" _ "\<lambda>x y. True", simplified]) auto
+
+lemmas alpha_bn_inducts = alpha_bn_inducts_raw[quot_lifted]
+
+nominal_primrec
+ subst :: "name \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm"
+and substa :: "name \<Rightarrow> trm \<Rightarrow> assn \<Rightarrow> assn"
+where
+ "subst s t (Var x) = (if (s = x) then t else (Var x))"
+| "subst s t (App l r) = App (subst s t l) (subst s t r)"
+| "atom v \<sharp> (s, t) \<Longrightarrow> subst s t (Lam v b) = Lam v (subst s t b)"
+| "set (bn as) \<sharp>* (s, t) \<Longrightarrow> subst s t (Let as b) = Let (substa s t as) (subst s t b)"
+| "substa s t ANil = ANil"
+| "substa s t (ACons v t' as) = ACons v (subst v t t') as"
+
+(*apply (subgoal_tac "\<forall>l. \<exists>!r. subst_substa_graph l r")
+defer
+apply rule
+apply (simp only: Ex1_def)
+apply (case_tac l)
+apply (case_tac a)
+apply (rule_tac y="c" and c="(aa,b)" in trm_assn.strong_exhaust(1))
+apply simp_all[3]
+apply rule
+apply rule
+apply (rule subst_substa_graph.intros)*)
+
+defer
+apply (case_tac x)
+apply (case_tac a)
+thm trm_assn.strong_exhaust(1)
+apply (rule_tac y="c" and c="(aa,b)" in trm_assn.strong_exhaust(1))
+apply (simp add: trm_assn.distinct trm_assn.eq_iff)
+apply auto[1]
+apply (simp add: trm_assn.distinct trm_assn.eq_iff)
+apply auto[1]
+apply (simp add: trm_assn.distinct trm_assn.eq_iff fresh_star_def)
+apply (simp add: trm_assn.distinct trm_assn.eq_iff)
+apply (drule_tac x="assn" in meta_spec)
+apply (rotate_tac 3)
+apply (drule_tac x="aa" in meta_spec)
+apply (rotate_tac -1)
+apply (drule_tac x="b" in meta_spec)
+apply (rotate_tac -1)
+apply (drule_tac x="trm" in meta_spec)
+apply (auto simp add: alpha_bn_refl)[1]
+apply (case_tac b)
+apply (rule_tac ya="c" in trm_assn.strong_exhaust(2))
+apply (simp add: trm_assn.distinct trm_assn.eq_iff)
+apply auto[1]
+apply blast
+apply (simp add: trm_assn.distinct trm_assn.eq_iff)
+apply auto[1]
+apply blast
+apply (simp_all only: sum.simps Pair_eq trm_assn.distinct trm_assn.eq_iff)
+apply simp_all
+apply (simp_all add: meta_eq_to_obj_eq[OF subst_def, symmetric, unfolded fun_eq_iff])
+apply (simp_all add: meta_eq_to_obj_eq[OF substa_def, symmetric, unfolded fun_eq_iff])
+apply clarify
+prefer 2
+apply clarify
+apply (rule conjI)
+prefer 2
+apply (rename_tac a pp vv zzz a2 s t zz)
+apply (erule alpha_bn_inducts)
+apply (rule alpha_bn_refl)
+apply clarify
+apply (rename_tac t' a1 a2 n1 n2)
+thm subst_substa_graph.intros[no_vars]
+.
+alpha_bn (substa s t (ACons n1 t' a1))
+ (substa s t (ACons n2 t' a2))
+
+alpha_bn (Acons s (subst a t t') a1)
+ (Acons s (subst a t t') a2)
+
+ACons v (subst v t t') as"
end