theory LetSimple1imports "../Nominal2" beginlemma Abs_lst_fcb2: fixes as bs :: "atom list" and x y :: "'b :: fs" and c::"'c::fs" assumes eq: "[as]lst. x = [bs]lst. y" and fcb1: "(set as) \<sharp>* f as x c" and fresh1: "set as \<sharp>* c" and fresh2: "set bs \<sharp>* c" and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c" and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c" shows "f as x c = f bs y c"proof - have "supp (as, x, c) supports (f as x c)" unfolding supports_def fresh_def[symmetric] by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh) then have fin1: "finite (supp (f as x c))" by (auto intro: supports_finite simp add: finite_supp) have "supp (bs, y, c) supports (f bs y c)" unfolding supports_def fresh_def[symmetric] by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh) then have fin2: "finite (supp (f bs y c))" by (auto intro: supports_finite simp add: finite_supp) obtain q::"perm" where fr1: "(q \<bullet> (set as)) \<sharp>* (x, c, f as x c, f bs y c)" and fr2: "supp q \<sharp>* Abs_lst as x" and inc: "supp q \<subseteq> (set as) \<union> q \<bullet> (set as)" using at_set_avoiding3[where xs="set as" and c="(x, c, f as x c, f bs y c)" and x="[as]lst. x"] fin1 fin2 by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv) have "Abs_lst (q \<bullet> as) (q \<bullet> x) = q \<bullet> Abs_lst as x" by simp also have "\<dots> = Abs_lst as x" by (simp only: fr2 perm_supp_eq) finally have "Abs_lst (q \<bullet> as) (q \<bullet> x) = Abs_lst bs y" using eq by simp then obtain r::perm where qq1: "q \<bullet> x = r \<bullet> y" and qq2: "q \<bullet> as = r \<bullet> bs" and qq3: "supp r \<subseteq> (q \<bullet> (set as)) \<union> set bs" apply(drule_tac sym) apply(simp only: Abs_eq_iff2 alphas) apply(erule exE) apply(erule conjE)+ apply(drule_tac x="p" in meta_spec) apply(simp add: set_eqvt) apply(blast) done have "(set as) \<sharp>* f as x c" by (rule fcb1) then have "q \<bullet> ((set as) \<sharp>* f as x c)" by (simp add: permute_bool_def) then have "set (q \<bullet> as) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c" apply(simp add: fresh_star_eqvt set_eqvt) apply(subst (asm) perm1) using inc fresh1 fr1 apply(auto simp add: fresh_star_def fresh_Pair) done then have "set (r \<bullet> bs) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp then have "r \<bullet> ((set bs) \<sharp>* f bs y c)" apply(simp add: fresh_star_eqvt set_eqvt) apply(subst (asm) perm2[symmetric]) using qq3 fresh2 fr1 apply(auto simp add: set_eqvt fresh_star_def fresh_Pair) done then have fcb2: "(set bs) \<sharp>* f bs y c" by (simp add: permute_bool_def) have "f as x c = q \<bullet> (f as x c)" apply(rule perm_supp_eq[symmetric]) using inc fcb1 fr1 by (auto simp add: fresh_star_def) also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c" apply(rule perm1) using inc fresh1 fr1 by (auto simp add: fresh_star_def) also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp also have "\<dots> = r \<bullet> (f bs y c)" apply(rule perm2[symmetric]) using qq3 fresh2 fr1 by (auto simp add: fresh_star_def) also have "... = f bs y c" apply(rule perm_supp_eq) using qq3 fr1 fcb2 by (auto simp add: fresh_star_def) finally show ?thesis by simpqedlemma Abs_lst1_fcb2: fixes a b :: "atom" and x y :: "'b :: fs" and c::"'c :: fs" assumes e: "(Abs_lst [a] x) = (Abs_lst [b] y)" and fcb1: "a \<sharp> f a x c" and fresh: "{a, b} \<sharp>* c" and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f a x c) = f (p \<bullet> a) (p \<bullet> x) c" and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f b y c) = f (p \<bullet> b) (p \<bullet> y) c" shows "f a x c = f b y c"using eapply(drule_tac Abs_lst_fcb2[where c="c" and f="\<lambda>(as::atom list) . f (hd as)"])apply(simp_all)using fcb1 fresh perm1 perm2apply(simp_all add: fresh_star_def)doneatom_decl namenominal_datatype trm = Var "name"| App "trm" "trm"| Let x::"name" "trm" t::"trm" bind x in tprint_theoremsthm trm.fv_defsthm trm.eq_iff thm trm.bn_defsthm trm.bn_inductsthm trm.perm_simpsthm trm.inductthm trm.inductsthm trm.distinctthm trm.suppthm trm.freshthm trm.exhaustthm trm.strong_exhaustthm trm.perm_bn_simpsnominal_primrec height_trm :: "trm \<Rightarrow> nat"where "height_trm (Var x) = 1"| "height_trm (App l r) = max (height_trm l) (height_trm r)"| "height_trm (Let x t s) = max (height_trm t) (height_trm s)" apply (simp only: eqvt_def height_trm_graph_def) apply (rule, perm_simp, rule, rule TrueI) apply (case_tac x rule: trm.exhaust(1)) apply (auto)[3] apply(simp_all)[5] apply (subgoal_tac "height_trm_sumC t = height_trm_sumC ta") apply (subgoal_tac "height_trm_sumC s = height_trm_sumC sa") apply simp apply(simp) apply(erule conjE) apply(erule_tac c="()" in Abs_lst1_fcb2) apply(simp_all add: fresh_star_def pure_fresh)[2] apply(simp_all add: eqvt_at_def)[2] apply(simp) donetermination by lexicographic_ordernominal_primrec (invariant "\<lambda>x (y::atom set). finite y") frees_set :: "trm \<Rightarrow> atom set"where "frees_set (Var x) = {atom x}"| "frees_set (App t1 t2) = frees_set t1 \<union> frees_set t2"| "frees_set (Let x t s) = (frees_set s) - {atom x} \<union> (frees_set t)" apply(simp add: eqvt_def frees_set_graph_def) apply(rule, perm_simp, rule) apply(erule frees_set_graph.induct) apply(auto)[3] apply(rule_tac y="x" in trm.exhaust) apply(auto)[3] apply(simp_all)[5] apply(simp) apply(erule conjE) apply(subgoal_tac "frees_set_sumC s - {atom x} = frees_set_sumC sa - {atom xa}") apply(simp) apply(erule_tac c="()" in Abs_lst1_fcb2) apply(simp add: fresh_minus_atom_set) apply(simp add: fresh_star_def fresh_Unit) apply(simp add: Diff_eqvt eqvt_at_def, perm_simp, rule refl) apply(simp add: Diff_eqvt eqvt_at_def, perm_simp, rule refl) donetermination by lexicographic_ordernominal_primrec subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_ [_ ::= _]" [90, 90, 90] 90)where "(Var x)[y ::= s] = (if x = y then s else (Var x))"| "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])"| "atom x \<sharp> (y, s) \<Longrightarrow> (Let x t t')[y ::= s] = Let x (t[y ::= s]) (t'[y ::= s])" apply(simp add: eqvt_def subst_graph_def) apply (rule, perm_simp, rule) apply(rule TrueI) apply(auto)[1] apply(rule_tac y="a" and c="(aa, b)" in trm.strong_exhaust) apply(blast)+ apply(simp_all add: fresh_star_def fresh_Pair_elim)[1] apply(blast) apply(simp_all)[5] apply(simp (no_asm_use)) apply(simp) apply(erule conjE)+ apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2) apply(simp add: Abs_fresh_iff) apply(simp add: fresh_star_def fresh_Pair) apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)donetermination by lexicographic_orderend