diff -r 6d9018d62b40 -r 679cef364022 Nominal/Ex/SingleLet.thy --- a/Nominal/Ex/SingleLet.thy Mon Oct 04 12:39:57 2010 +0100 +++ b/Nominal/Ex/SingleLet.thy Tue Oct 05 07:30:37 2010 +0100 @@ -21,12 +21,10 @@ where "bn (As x y t) = [atom x]" - - thm single_let.distinct thm single_let.induct thm single_let.inducts -thm single_let.exhaust +thm single_let.exhaust[no_vars] thm single_let.fv_defs thm single_let.bn_defs thm single_let.perm_simps @@ -38,6 +36,131 @@ thm single_let.supp thm single_let.fresh +primrec + permute_bn_raw +where + "permute_bn_raw p (As_raw x y t) = As_raw (p \ x) y t" + +quotient_definition + "permute_bn :: perm \ assg \ assg" +is + "permute_bn_raw" + +lemma [quot_respect]: + shows "((op =) ===> alpha_assg_raw ===> alpha_assg_raw) permute_bn_raw permute_bn_raw" + apply simp + apply clarify + apply (erule alpha_trm_raw_alpha_assg_raw_alpha_bn_raw.inducts) + apply (rule TrueI)+ + apply simp_all + apply (rule_tac [!] alpha_trm_raw_alpha_assg_raw_alpha_bn_raw.intros) + apply simp_all + done + +lemmas permute_bn = permute_bn_raw.simps[quot_lifted] + +lemma uu: + shows "alpha_bn as (permute_bn p as)" +apply(induct as rule: single_let.inducts(2)) +apply(auto)[7] +apply(simp add: permute_bn) +apply(simp add: single_let.eq_iff) +done + +lemma tt: + shows "(p \ bn as) = bn (permute_bn p as)" +apply(induct as rule: single_let.inducts(2)) +apply(auto)[7] +apply(simp add: permute_bn single_let.bn_defs) +apply(simp add: atom_eqvt) +done + +lemma Abs_fresh_star': + fixes x::"'a::fs" + shows "set bs = as \ as \* Abs_lst bs x" + unfolding fresh_star_def + by(simp_all add: Abs_fresh_iff) + +lemma strong_exhaust: + fixes c::"'a::fs" + assumes "\name. y = Var name \ P" + and "\trm1 trm2. y = App trm1 trm2 \ P" + and "\name trm. \{atom name} \* c; y = Lam name trm\ \ P" + and "\assg trm. \set (bn assg) \* c; y = Let assg trm\ \ P" + and "\name1 name2 trm1 trm2 trm3. \{atom name1} \* c; y = Foo name1 name2 trm1 trm2 trm3\ \ P" + and "\name1 name2 trm. \{atom name1, atom name2} \* c; y = Bar name1 name2 trm\ \ P" + and "\name trm1 trm2. \{atom name} \* c; y = Baz name trm1 trm2\ \ P" + shows "P" + apply(rule_tac y="y" in single_let.exhaust(1)) + apply(rule assms(1)) + apply(assumption) + apply(rule assms(2)) + apply(assumption) + apply(subgoal_tac "\q. (q \ {atom name}) \* c \ supp ([[atom name]]lst.trm) \* q") + apply(erule exE) + apply(erule conjE) + apply(perm_simp) + apply(rule assms(3)) + apply(assumption) + apply(drule supp_perm_eq[symmetric]) + apply(simp add: single_let.eq_iff) + apply(perm_simp) + apply(rule refl) + apply(rule at_set_avoiding2) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: Abs_fresh_star') + apply(subgoal_tac "\q. (q \ (set (bn assg))) \* (c::'a::fs) \ supp ([bn assg]lst.trm) \* q") + apply(erule exE) + apply(erule conjE) + apply(perm_simp add: tt) + apply(rule_tac assms(4)) + apply(assumption) + apply(drule supp_perm_eq[symmetric]) + apply(simp add: single_let.eq_iff) + apply(simp add: tt uu) + apply(rule at_set_avoiding2) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: Abs_fresh_star) + apply(subgoal_tac "\q. (q \ {atom name1}) \* (c::'a::fs) \ + supp ([{atom name1}]set.(((name2, trm1), trm2), trm3)) \* q") + apply(erule exE) + apply(erule conjE) + apply(perm_simp add: tt) + apply(rule_tac assms(5)) + apply(assumption) + apply(drule supp_perm_eq[symmetric]) + apply(simp add: single_let.eq_iff) + apply(perm_simp) + apply(rule refl) + apply(rule at_set_avoiding2) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: Abs_fresh_star) + apply(subgoal_tac "\q. (q \ {atom name1, atom name2}) \* (c::'a::fs) \ + supp ([[atom name2, atom name1]]lst.((trm, name1), name2)) \* q") + apply(erule exE) + apply(erule conjE) + apply(perm_simp add: tt) + apply(rule_tac assms(6)) + apply(assumption) + apply(drule supp_perm_eq[symmetric]) + apply(simp add: single_let.eq_iff) + apply(perm_simp) + apply(rule refl) + apply(rule at_set_avoiding2) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + apply(simp add: finite_supp) + oops + (*apply(simp add: Abs_fresh_star)*) +sorry + +done end