--- a/Nominal/Ex/SingleLet.thy Fri Sep 03 20:53:09 2010 +0800
+++ b/Nominal/Ex/SingleLet.thy Fri Sep 03 22:22:43 2010 +0800
@@ -9,17 +9,17 @@
nominal_datatype single_let: trm =
Var "name"
| App "trm" "trm"
-| Lam x::"name" t::"trm" bind (set) x in t
-| Let a::"assg" t::"trm" bind (set) "bn a" in t
+| Lam x::"name" t::"trm" bind x in t
+| Let a::"assg" t::"trm" bind "bn a" in t
| Foo x::"name" y::"name" t::"trm" t1::"trm" t2::"trm" bind (set) x in y t t1 t2
| Bar x::"name" y::"name" t::"trm" bind y x in t x y
| Baz x::"name" t1::"trm" t2::"trm" bind x in t1, bind x in t2
and assg =
As "name" x::"name" t::"trm" bind x in t
binder
- bn::"assg \<Rightarrow> atom set"
+ bn::"assg \<Rightarrow> atom list"
where
- "bn (As x y t) = {atom x}"
+ "bn (As x y t) = [atom x]"
thm single_let.distinct
@@ -34,20 +34,6 @@
thm single_let.supports
thm single_let.fsupp
-lemma supp_abs_sum:
- fixes a b::"'a::fs"
- shows "supp (Abs x a) \<union> supp (Abs x b) = supp (Abs x (a, b))"
- and "supp (Abs_res x a) \<union> supp (Abs_res x b) = supp (Abs_res x (a, b))"
- and "supp (Abs_lst y a) \<union> supp (Abs_lst y b) = supp (Abs_lst y (a, b))"
- apply (simp_all add: supp_abs supp_Pair)
- apply blast+
- done
-
-
-lemma test:
- "(\<exists>p. (bs, x) \<approx>gen (op=) f p (cs, y)) \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op=) supp p (cs, y))"
-sorry
-
lemma Abs_eq_iff:
shows "Abs bs x = Abs cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op =) supp p (cs, y))"
and "Abs_res bs x = Abs_res cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>res (op =) supp p (cs, y))"
@@ -57,126 +43,96 @@
lemma test2:
assumes "fv_trm t = supp t"
shows "\<forall>p. fv_trm (p \<bullet> t) = supp (p \<bullet> t)"
-sorry
-
-lemma yy:
- "X = Y \<Longrightarrow> finite X = finite Y" by simp
+apply(rule allI)
+apply(rule_tac p="-p" in permute_boolE)
+apply(perm_simp add: single_let.fv_bn_eqvt permute_minus_cancel)
+apply(rule assms)
+done
lemma supp_fv:
"fv_trm t = supp t \<and> fv_assg as = supp as \<and> fv_bn as = {a. infinite {b. \<not>alpha_bn ((a \<rightleftharpoons> b) \<bullet> as) as}}"
apply(rule single_let.induct)
-apply(simp_all only: single_let.fv_defs)[2]
-apply(simp_all only: supp_def)[2]
-apply(simp_all only: single_let.perm_simps)[2]
-apply(simp_all only: single_let.eq_iff)[2]
-apply(simp_all only: de_Morgan_conj)[2]
-apply(simp_all only: Collect_disj_eq)[2]
-apply(simp_all only: finite_Un)[2]
-apply(simp_all only: de_Morgan_conj)[2]
-apply(simp_all only: Collect_disj_eq)[2]
+apply(simp_all (no_asm_use) only: single_let.fv_defs)[2]
+apply(simp_all (no_asm_use) only: supp_def)[2]
+apply(simp_all (no_asm_use) only: single_let.perm_simps)[2]
+apply(simp_all (no_asm_use) only: single_let.eq_iff)[2]
+apply(simp_all (no_asm_use) only: de_Morgan_conj)[2]
+apply(simp_all (no_asm_use) only: Collect_disj_eq)[2]
+apply(simp_all (no_asm_use) only: finite_Un)[2]
+apply(simp_all (no_asm_use) only: de_Morgan_conj)[2]
+apply(simp_all (no_asm_use) only: Collect_disj_eq)[2]
+apply(simp)
--" 1 "
-apply(simp only: single_let.fv_defs)
-apply(simp add: supp_abs(1)[symmetric])
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp)
-apply(simp only: Abs_eq_iff)
-apply(simp add: alphas)
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
apply(drule test2)
-apply(simp)
+apply(simp only:)
-- " 2 "
apply(erule conjE)+
-apply(simp only: single_let.fv_defs)
-apply(simp add: supp_abs(1)[symmetric])
-apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: de_Morgan_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp only: Abs_eq_iff)
-apply(simp only: finite_Un)
-apply(simp only: de_Morgan_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp add: alphas)
-apply(drule test2)
-apply(simp)
--- " 3 "
-apply(simp only: single_let.fv_defs)
-apply(simp only: supp_Pair[symmetric])
-apply(simp add: supp_abs(1)[symmetric])
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: Abs_eq_iff)
-apply(simp add: alphas)
-apply(simp add: supp_Pair)
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(drule test2)
+apply(simp only:)
+-- " 3 "
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(1)[symmetric])
+apply(simp (no_asm) only: supp_def)
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
apply(drule test2)+
-apply(simp)
-apply(simp add: prod_alpha_def)
-apply(simp add: Un_assoc)
-apply(rule Collect_cong)
-apply(rule arg_cong)
-back
-apply(rule yy)
-apply(rule Collect_cong)
-apply(auto)[1]
+apply(simp only: supp_Pair Un_assoc conj_assoc)
-- " Bar "
-apply(simp only: single_let.fv_defs)
-apply(simp only: supp_Pair[symmetric])
-apply(simp add: supp_abs(1)[symmetric])
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: Abs_eq_iff)
-apply(simp add: alphas prod_alpha_def)
-apply(drule test2)
-apply(simp add: supp_Pair)
-apply(subst atom_eqvt)
-apply(simp)
-apply(simp add: Un_assoc)
-apply(rule Collect_cong)
-apply(rule arg_cong)
-back
-apply(rule yy)
-apply(rule Collect_cong)
--- "last"
-prefer 3
-apply(rule conjI)
-apply(simp only: single_let.fv_defs)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp add: supp_abs(1)[symmetric])
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(drule test2)+
+apply(simp only: supp_Pair Un_assoc conj_assoc)
+-- " Baz "
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
apply(simp (no_asm) only: supp_def)
-apply(simp only: single_let.perm_simps)
-apply(simp only: single_let.eq_iff)
-apply(simp only: permute_abs atom_eqvt permute_list.simps)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: Abs_eq_iff)
-apply(simp only: de_Morgan_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp only: finite_Un)
-apply(simp only: de_Morgan_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp add: alphas prod_alpha_def)
-apply(drule test2)
-apply(simp add: supp_Pair)
-apply(simp only: permute_set_eq)
-apply(simp)
-apply(perm_simp add: single_let.fv_bn_eqvt)
-apply(simp only: single_let.eq_iff)
-apply(simp only: single_let.fv_defs)
-apply(simp add: supp_abs(1)[symmetric])
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)
+apply(drule test2)+
+apply(simp only: supp_Pair Un_assoc conj_assoc)
+-- "last"
+apply(rule conjI)
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
apply(simp (no_asm) only: supp_def)
-apply(perm_simp)
-oops
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)
+apply(drule test2)+
+apply(simp only: supp_Pair Un_assoc conj_assoc)
+-- "other case"
+apply(simp only: single_let.fv_defs supp_Pair[symmetric])
+apply(simp only: supp_abs(3)[symmetric])
+apply(simp (no_asm) only: supp_def)
+apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt)
+apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff)
+apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject)
+apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)?
+apply(drule test2)+
+apply(simp only: supp_Pair Un_assoc conj_assoc)
+done