diff -r 1d9e50934bc5 -r aaef9dec5e1d Nominal/Ex/Let.thy --- a/Nominal/Ex/Let.thy Fri Jul 01 17:46:15 2011 +0900 +++ b/Nominal/Ex/Let.thy Sat Jul 02 00:27:47 2011 +0100 @@ -39,8 +39,9 @@ thm trm_assn.fresh thm trm_assn.exhaust thm trm_assn.strong_exhaust +thm trm_assn.perm_bn_simps -lemma alpha_bn_inducts_raw: +lemma alpha_bn_inducts_raw[consumes 1]: "\alpha_bn_raw a b; P3 ANil_raw ANil_raw; \trm_raw trm_rawa assn_raw assn_rawa name namea. \alpha_trm_raw trm_raw trm_rawa; alpha_bn_raw assn_raw assn_rawa; @@ -49,7 +50,7 @@ (ACons_raw namea trm_rawa assn_rawa)\ \ P3 a b" by (erule alpha_trm_raw_alpha_assn_raw_alpha_bn_raw.inducts(3)[of _ _ "\x y. True" _ "\x y. True", simplified]) auto -lemmas alpha_bn_inducts = alpha_bn_inducts_raw[quot_lifted] +lemmas alpha_bn_inducts[consumes 1] = alpha_bn_inducts_raw[quot_lifted] @@ -66,6 +67,17 @@ shows "bn x = bn y \ x = y" by (rule alpha_bn_inducts[OF a]) (simp_all add: trm_assn.bn_defs) +lemma bn_inj2: + assumes a: "alpha_bn x y" + shows "\q r. (q \ bn x) = (r \ bn y) \ permute_bn q x = permute_bn r y" +using a +apply(induct rule: alpha_bn_inducts) +apply(simp add: trm_assn.perm_bn_simps) +apply(simp add: trm_assn.perm_bn_simps) +apply(simp add: trm_assn.bn_defs) +apply(simp add: atom_eqvt) +done + (*lemma alpha_bn_permute: assumes a: "alpha_bn x y" and b: "q \ bn x = r \ bn y" @@ -85,6 +97,205 @@ using bn_inj by simp *) + +lemma max_eqvt[eqvt]: "p \ (max (a :: _ :: pure) b) = max (p \ a) (p \ b)" + by (simp add: permute_pure) + + +lemma Abs_lst_fcb2: + fixes as bs :: "'a :: fs" + and x y :: "'b :: fs" + assumes eq: "[ba as]lst. x = [ba bs]lst. y" + and ctxt: "finite (supp c)" + and fcb1: "set (ba as) \* f as x c" + and fresh1: "set (ba as) \* c" + and fresh2: "set (ba 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" +(* What we would like in this proof, and lets this proof finish *) + and ba_inj: "\q r. q \ ba as = r \ ba bs \ pn q as = pn r bs" + 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] + apply (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh) + sorry + then have fin1: "finite (supp (f as x c))" + by (auto intro: supports_finite simp add: finite_supp supp_Pair ctxt) + have "supp (bs, y, c) supports (f bs y c)" + unfolding supports_def fresh_def[symmetric] + apply (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh) + sorry + then have fin2: "finite (supp (f bs y c))" + by (auto intro: supports_finite simp add: finite_supp supp_Pair ctxt) + obtain q::"perm" where + fr1: "(q \ (set (ba as))) \* (x, c, f as x c, f bs y c)" and + fr2: "supp q \* ([ba as]lst. x)" and + inc: "supp q \ (set (ba as)) \ q \ (set (ba as))" + using at_set_avoiding3[where xs="set (ba as)" and c="(x, c, f as x c, f bs y c)" + and x="[ba as]lst. x"] fin1 fin2 + by (auto simp add: supp_Pair finite_supp ctxt Abs_fresh_star dest: fresh_star_supp_conv) + have "[q \ (ba as)]lst. (q \ x) = q \ ([ba as]lst. x)" by simp + also have "\ = [ba as]lst. x" + by (simp only: fr2 perm_supp_eq) + finally have "[q \ (ba as)]lst. (q \ x) = [ba bs]lst. y" using eq by simp + then obtain r::perm where + qq1: "q \ x = r \ y" and + qq2: "q \ (ba as) = r \ (ba bs)" and + qq3: "supp r \ (q \ (set (ba as))) \ set (ba 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 (ba as)) \* f as x c" by (rule fcb1) + then have "q \ ((set (ba as)) \* f as x c)" + by (simp add: permute_bool_def) + then have "set (q \ (ba 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 \ (ba bs)) \* f (r \ bs) (r \ y) c" using qq1 qq2 ba_inj + apply simp + sorry + then have "r \ ((set (ba 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 (ba 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 ba_inj + apply(simp) + sorry + 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 + +nominal_primrec + height_trm :: "trm \ nat" +and height_assn :: "assn \ 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 as b) = max (height_assn as) (height_trm b)" +| "height_assn ANil = 0" +| "height_assn (ACons v t as) = max (height_trm t) (height_assn as)" + apply (simp only: eqvt_def height_trm_height_assn_graph_def) + apply (rule, perm_simp, rule, rule TrueI) + apply (case_tac x) + apply (case_tac a rule: trm_assn.exhaust(1)) + apply (auto)[4] + apply (drule_tac x="assn" in meta_spec) + apply (drule_tac x="trm" in meta_spec) + apply (simp add: alpha_bn_refl) + apply (case_tac b rule: trm_assn.exhaust(2)) + apply (auto)[2] + apply(simp_all del: trm_assn.eq_iff) + apply(simp) + prefer 3 + apply(simp) + apply(simp) + apply (erule_tac c="()" and pn="permute" in Abs_lst_fcb2) + apply(simp add: finite_supp) + apply (simp_all add: pure_fresh fresh_star_def)[3] + apply (simp add: eqvt_at_def) + apply (simp add: eqvt_at_def) + apply(auto)[1] + --"other case" + apply (simp only: meta_eq_to_obj_eq[OF height_trm_def, symmetric, unfolded fun_eq_iff]) + apply (simp only: meta_eq_to_obj_eq[OF height_assn_def, symmetric, unfolded fun_eq_iff]) + apply (subgoal_tac "eqvt_at height_assn as") + apply (subgoal_tac "eqvt_at height_assn asa") + apply (subgoal_tac "eqvt_at height_trm b") + apply (subgoal_tac "eqvt_at height_trm ba") + apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inr as)") + apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inr asa)") + apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inl b)") + apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inl ba)") + defer + apply (simp add: eqvt_at_def height_trm_def) + apply (simp add: eqvt_at_def height_trm_def) + apply (simp add: eqvt_at_def height_assn_def) + apply (simp add: eqvt_at_def height_assn_def) + apply (subgoal_tac "height_assn as = height_assn asa") + apply (subgoal_tac "height_trm b = height_trm ba") + apply simp + apply(simp) + apply(erule conjE) + apply (erule_tac c="()" and pn="permute_bn" in Abs_lst_fcb2) + apply(simp add: finite_supp) + apply (simp_all add: pure_fresh fresh_star_def)[3] + apply (simp_all add: eqvt_at_def)[2] + apply(simp add: bn_inj2) + apply(simp) + apply(erule conjE) + thm trm_assn.fv_defs + (*apply(simp add: Abs_eq_iff alphas)*) + apply (erule_tac c="()" and pn="permute_bn" and ba="bn" in Abs_lst_fcb2) + defer + apply (simp_all add: pure_fresh fresh_star_def)[3] + defer + defer + apply (simp_all add: eqvt_at_def)[2] + apply (rule bn_inj) + prefer 2 + apply (simp add: eqvts) + oops + +nominal_primrec + subst :: "name \ trm \ trm \ trm" +and substa :: "name \ trm \ assn \ 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 \ (s, t) \ subst s t (Lam v b) = Lam v (subst s t b)" +| "set (bn as) \* (s, t) \ 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" +(*unfolding eqvt_def subst_substa_graph_def + apply (rule, perm_simp)*) + defer + apply (rule TrueI) + apply (case_tac x) + apply (case_tac a) + apply (rule_tac y="c" and c="(aa,b)" in trm_assn.strong_exhaust(1)) + apply (auto simp add: fresh_star_def)[3] + apply (drule_tac x="assn" in meta_spec) + apply (simp add: Abs1_eq_iff alpha_bn_refl) + apply (case_tac b) + apply (case_tac c rule: trm_assn.exhaust(2)) + apply (auto)[2] + apply blast + apply blast + apply auto + 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]) + prefer 2 + apply (erule_tac c="(sa, ta)" in Abs_lst_fcb2) + apply (simp_all add: fresh_star_Pair) + prefer 6 + apply (erule alpha_bn_inducts) + oops + + lemma lets_bla: "x \ z \ y \ z \ x \ y \(Let (ACons x (Var y) ANil) (Var x)) \ (Let (ACons x (Var z) ANil) (Var x))" by (simp add: trm_assn.eq_iff) @@ -126,188 +337,4 @@ 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 max_eqvt[eqvt]: "p \ (max (a :: _ :: pure) b) = max (p \ a) (p \ b)" - by (simp add: permute_pure) - - -lemma Abs_lst_fcb2: - fixes as bs :: "'a :: fs" - and x y :: "'b :: fs" - and c::"'c::fs" - assumes eq: "[ba as]lst. x = [ba bs]lst. y" - and fcb1: "set (ba as) \* f as x c" - and fresh1: "set (ba as) \* c" - and fresh2: "set (ba 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" -(* What we would like in this proof, and lets this proof finish *) - and ba_inj: "\q r. q \ ba as = r \ ba bs \ q \ as = r \ bs" -(* What the user can supply with the help of alpha_bn *) -(* and bainj: "ba as = ba bs \ as = bs"*) - 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 (ba as))) \* (x, c, f as x c, f bs y c)" and - fr2: "supp q \* ([ba as]lst. x)" and - inc: "supp q \ (set (ba as)) \ q \ (set (ba as))" - using at_set_avoiding3[where xs="set (ba as)" and c="(x, c, f as x c, f bs y c)" - and x="[ba as]lst. x"] fin1 fin2 - by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv) - have "[q \ (ba as)]lst. (q \ x) = q \ ([ba as]lst. x)" by simp - also have "\ = [ba as]lst. x" - by (simp only: fr2 perm_supp_eq) - finally have "[q \ (ba as)]lst. (q \ x) = [ba bs]lst. y" using eq by simp - then obtain r::perm where - qq1: "q \ x = r \ y" and - qq2: "q \ (ba as) = r \ (ba bs)" and - qq3: "supp r \ (q \ (set (ba as))) \ set (ba 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 (ba as)) \* f as x c" by (rule fcb1) - then have "q \ ((set (ba as)) \* f as x c)" - by (simp add: permute_bool_def) - then have "set (q \ (ba 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 \ (ba bs)) \* f (r \ bs) (r \ y) c" using qq1 qq2 ba_inj - by simp - then have "r \ ((set (ba 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 (ba 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 ba_inj 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 - -nominal_primrec - height_trm :: "trm \ nat" -and height_assn :: "assn \ 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 as b) = max (height_assn as) (height_trm b)" -| "height_assn ANil = 0" -| "height_assn (ACons v t as) = max (height_trm t) (height_assn as)" - apply (simp only: eqvt_def height_trm_height_assn_graph_def) - apply (rule, perm_simp, rule, rule TrueI) - apply (case_tac x) - apply (case_tac a rule: trm_assn.exhaust(1)) - apply (auto)[4] - apply (drule_tac x="assn" in meta_spec) - apply (drule_tac x="trm" in meta_spec) - apply (simp add: alpha_bn_refl) - apply (case_tac b rule: trm_assn.exhaust(2)) - apply (auto)[2] - apply(simp_all) - apply (erule_tac c="()" in Abs_lst_fcb2) - apply (simp_all add: pure_fresh fresh_star_def)[3] - apply (simp add: eqvt_at_def) - apply (simp add: eqvt_at_def) - apply assumption - apply(erule conjE) - apply (simp add: meta_eq_to_obj_eq[OF height_trm_def, symmetric, unfolded fun_eq_iff]) - apply (simp add: meta_eq_to_obj_eq[OF height_assn_def, symmetric, unfolded fun_eq_iff]) - apply (subgoal_tac "eqvt_at height_assn as") - apply (subgoal_tac "eqvt_at height_assn asa") - apply (subgoal_tac "eqvt_at height_trm b") - apply (subgoal_tac "eqvt_at height_trm ba") - apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inr as)") - apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inr asa)") - apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inl b)") - apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inl ba)") - defer - apply (simp add: eqvt_at_def height_trm_def) - apply (simp add: eqvt_at_def height_trm_def) - apply (simp add: eqvt_at_def height_assn_def) - apply (simp add: eqvt_at_def height_assn_def) - apply (subgoal_tac "height_assn as = height_assn asa") - apply (subgoal_tac "height_trm b = height_trm ba") - apply simp - apply (erule_tac c="()" in Abs_lst_fcb2) - apply (simp_all add: pure_fresh fresh_star_def)[3] - apply (simp_all add: eqvt_at_def)[2] - apply assumption - apply (erule_tac c="()" and ba="bn" in Abs_lst_fcb2) - apply (simp_all add: pure_fresh fresh_star_def)[3] - apply (simp_all add: eqvt_at_def)[2] - apply (rule bn_inj) - prefer 2 - apply (simp add: eqvts) - oops - -nominal_primrec - subst :: "name \ trm \ trm \ trm" -and substa :: "name \ trm \ assn \ 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 \ (s, t) \ subst s t (Lam v b) = Lam v (subst s t b)" -| "set (bn as) \* (s, t) \ 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" -(*unfolding eqvt_def subst_substa_graph_def - apply (rule, perm_simp)*) - defer - apply (rule TrueI) - apply (case_tac x) - apply (case_tac a) - apply (rule_tac y="c" and c="(aa,b)" in trm_assn.strong_exhaust(1)) - apply (auto simp add: fresh_star_def)[3] - apply (drule_tac x="assn" in meta_spec) - apply (simp add: Abs1_eq_iff alpha_bn_refl) - apply (case_tac b) - apply (case_tac c rule: trm_assn.exhaust(2)) - apply (auto)[2] - apply blast - apply blast - apply auto - 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]) - prefer 2 - apply (erule_tac c="(sa, ta)" in Abs_lst_fcb2) - apply (simp_all add: fresh_star_Pair) - prefer 6 - apply (erule alpha_bn_inducts) - oops - - -end - - - +end \ No newline at end of file