diff -r db0786a521fd -r b4bf3ff4bc91 Nominal/Ex/LetRecB.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/Ex/LetRecB.thy Tue Jun 28 14:01:52 2011 +0100 @@ -0,0 +1,160 @@ +theory LetRecB +imports "../Nominal2" +begin + +atom_decl name + +nominal_datatype let_rec: + trm = + Var "name" +| App "trm" "trm" +| Lam x::"name" t::"trm" bind x in t +| Let_Rec bp::"bp" t::"trm" bind "bn bp" in bp t +and bp = + Bp "name" "trm" +binder + bn::"bp \ atom list" +where + "bn (Bp x t) = [atom x]" + +thm let_rec.distinct +thm let_rec.induct +thm let_rec.exhaust +thm let_rec.fv_defs +thm let_rec.bn_defs +thm let_rec.perm_simps +thm let_rec.eq_iff +thm let_rec.fv_bn_eqvt +thm let_rec.size_eqvt + + +lemma 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) \* f as x c" + and fresh1: "set as \* c" + and fresh2: "set bs \* c" + and perm1: "\p. supp p \* c \ p \ (f as x c) = f (p \ as) (p \ x) c" + and perm2: "\p. supp p \* c \ p \ (f bs y c) = f (p \ bs) (p \ 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 \ (set as)) \* (x, c, f as x c, f bs y c)" and + fr2: "supp q \* Abs_lst as x" and + inc: "supp q \ (set as) \ q \ (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 \ as) (q \ x) = q \ Abs_lst as x" by simp + also have "\ = Abs_lst as x" + by (simp only: fr2 perm_supp_eq) + finally have "Abs_lst (q \ as) (q \ x) = Abs_lst bs y" using eq by simp + then obtain r::perm where + qq1: "q \ x = r \ y" and + qq2: "q \ as = r \ bs" and + qq3: "supp r \ (q \ (set as)) \ 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) \* f as x c" by (rule fcb1) + then have "q \ ((set as) \* f as x c)" + by (simp add: permute_bool_def) + then have "set (q \ as) \* f (q \ as) (q \ 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 \ bs) \* f (r \ bs) (r \ y) c" using qq1 qq2 by simp + then have "r \ ((set bs) \* 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) \* f bs y c" by (simp add: permute_bool_def) + have "f as x c = q \ (f as x c)" + apply(rule perm_supp_eq[symmetric]) + using inc fcb1 fr1 by (auto simp add: fresh_star_def) + also have "\ = f (q \ as) (q \ x) c" + apply(rule perm1) + using inc fresh1 fr1 by (auto simp add: fresh_star_def) + also have "\ = f (r \ bs) (r \ y) c" using qq1 qq2 by simp + also have "\ = r \ (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 simp +qed + + +lemma max_eqvt[eqvt]: "p \ (max (a :: _ :: pure) b) = max (p \ a) (p \ b)" + by (simp add: permute_pure) + +nominal_primrec + height_trm :: "trm \ nat" +and height_bp :: "bp \ nat" +where + "height_trm (Var x) = 1" +| "height_trm (App l r) = max (height_trm l) (height_trm r)" +| "height_trm (Lam v b) = 1 + (height_trm b)" +| "height_trm (Let_Rec bp b) = max (height_bp bp) (height_trm b)" +| "height_bp (Bp v t) = height_trm t" + --"eqvt" + apply (simp only: eqvt_def height_trm_height_bp_graph_def) + apply (rule, perm_simp, rule, rule TrueI) + --"completeness" + apply (case_tac x) + apply (case_tac a rule: let_rec.exhaust(1)) + apply (auto)[4] + apply (case_tac b rule: let_rec.exhaust(2)) + apply blast + apply(simp_all) + apply (erule_tac c="()" in Abs_lst_fcb2) + apply (simp_all add: fresh_star_def pure_fresh)[3] + apply (simp add: eqvt_at_def) + apply (simp add: eqvt_at_def) + --"HERE" + thm Abs_lst_fcb2 + apply(rule Abs_lst_fcb2) + --" does not fit the assumption " + + apply (drule_tac c="()" in Abs_lst_fcb2) + prefer 6 + apply(assumption) + apply (drule_tac c="()" in Abs_lst_fcb2) + apply (simp add: Abs_eq_iff2) + apply (simp add: alphas) + apply clarify + apply (rule trans) + apply(rule_tac p="p" in supp_perm_eq[symmetric]) + apply (simp add: pure_supp fresh_star_def) + apply (simp only: eqvts) + apply (simp add: eqvt_at_def) + done + +termination by lexicographic_order + +end + + +