# HG changeset patch # User Christian Urban # Date 1290647904 0 # Node ID 6c4372a1f2204be90135bfd2bcce98bb256a3b4f # Parent 3696659358c884e87cc69cec2bb6844514b343fe# Parent d8a676a6904795f10000b4d055d6991a6553e4d8 merged diff -r 3696659358c8 -r 6c4372a1f220 Nominal/Ex/Foo2.thy --- a/Nominal/Ex/Foo2.thy Wed Nov 24 02:36:21 2010 +0000 +++ b/Nominal/Ex/Foo2.thy Thu Nov 25 01:18:24 2010 +0000 @@ -34,7 +34,7 @@ thm foo.fv_defs thm foo.bn_defs thm foo.perm_simps -thm foo.eq_iff +thm foo.eq_iff(5) thm foo.fv_bn_eqvt thm foo.size_eqvt thm foo.supports @@ -75,6 +75,43 @@ apply(simp add: uu1) done +lemma Let2_rename: + assumes "(supp ([[atom x, atom y]]lst. t1)) \* p" and "(supp ([[atom y]]lst. t2)) \* p" + shows "Let2 x y t1 t2 = Let2 (p \ x) (p \ y) (p \ t1) (p \ t2)" + using assms + apply - + apply(drule supp_perm_eq[symmetric]) + apply(drule supp_perm_eq[symmetric]) + apply(simp only: foo.eq_iff) + apply(simp only: eqvts) + apply simp + done + +lemma Let2_rename2: + assumes "(supp ([[atom x, atom y]]lst. t1)) \* p" and "(atom y) \ p" + shows "Let2 x y t1 t2 = Let2 (p \ x) y (p \ t1) t2" + using assms + apply - + apply(drule supp_perm_eq[symmetric]) + apply(simp only: foo.eq_iff) + apply(simp only: eqvts) + apply simp + by (metis assms(2) atom_eqvt fresh_perm) + +lemma Let2_rename3: + assumes "(supp ([[atom x, atom y]]lst. t1)) \* p" + and "(supp ([[atom y]]lst. t2)) \* p" + and "(atom x) \ p" + shows "Let2 x y t1 t2 = Let2 x (p \ y) (p \ t1) (p \ t2)" + using assms + apply - + apply(drule supp_perm_eq[symmetric]) + apply(drule supp_perm_eq[symmetric]) + apply(simp only: foo.eq_iff) + apply(simp only: eqvts) + apply simp + by (metis assms(2) atom_eqvt fresh_perm) + lemma strong_exhaust1: fixes c::"'a::fs" assumes "\name. y = Var name \ P" @@ -82,7 +119,7 @@ and "\name trm. \{atom name} \* c; y = Lam name trm\ \ P" and "\assn1 trm1 assn2 trm2. \((set (bn assn1)) \ (set (bn assn2))) \* c; y = Let1 assn1 trm1 assn2 trm2\ \ P" - and "\x1 x2 trm1 trm2. \{atom x1, atom x2} \* c; y = Let2 x1 x2 trm1 trm2\ \ P" + and "\x1 x2 trm1 trm2. \{atom x1} \* c; y = Let2 x1 x2 trm1 trm2\ \ P" shows "P" apply(rule_tac y="y" in foo.exhaust(1)) apply(rule assms(1)) @@ -130,27 +167,109 @@ apply(simp add: finite_supp) apply(simp add: finite_supp) apply(simp add: Abs_fresh_star) -thm foo.eq_iff +apply(case_tac "name1 = name2") apply(subgoal_tac - "\q. (q \ {atom name1}) \* c \ supp ([[atom name1]]lst. trm1) \* q") -apply(subgoal_tac - "\q. (q \ {atom name2}) \* c \ supp ([[atom name2]]lst. trm2) \* q") + "\q. (q \ {atom name1, atom name2}) \* c \ (supp (([[atom name1, atom name2]]lst. trm1), ([[atom name2]]lst. trm2))) \* q") +apply(erule exE)+ +apply(erule conjE)+ +apply(perm_simp) +apply(rule assms(5)) + +apply (simp add: fresh_star_def eqvts) +apply (simp only: supp_Pair) +apply (simp only: fresh_star_Un_elim) +apply (subst Let2_rename) +apply assumption +apply assumption +apply (rule refl) +apply(rule at_set_avoiding2) +apply(simp add: finite_supp) +apply(simp add: finite_supp) +apply(simp add: finite_supp) +apply clarify +apply (simp add: fresh_star_def) +apply (simp add: fresh_def supp_Pair supp_Abs) +apply(subgoal_tac + "\q. (q \ {atom name1}) \* c \ (supp ((([[atom name1, atom name2]]lst. trm1)), (atom name2))) \* q") +prefer 2 +apply(rule at_set_avoiding2) +apply(simp add: finite_supp) +apply(simp add: finite_supp) +apply(simp add: finite_supp) +apply (simp add: fresh_star_def) +apply (simp add: fresh_def supp_Pair supp_Abs supp_atom) apply(erule exE)+ apply(erule conjE)+ +apply(perm_simp) apply(rule assms(5)) +apply assumption +apply clarify +apply (rule_tac x="name1" and y="name2" and ?t1.0="trm1" and ?t2.0="trm2" in Let2_rename2) +apply (simp_all add: fresh_star_Un_elim supp_Pair supp_Abs) +apply (simp add: fresh_star_def supp_atom) +done + +lemma strong_exhaust2: + 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 "\assn1 trm1 assn2 trm2. + \((set (bn assn1)) \ (set (bn assn2))) \* c; y = Let1 assn1 trm1 assn2 trm2\ \ P" + and "\x1 x2 trm1 trm2. \{atom x1, atom x2} \* c; y = Let2 x1 x2 trm1 trm2\ \ P" + shows "P" + apply (rule strong_exhaust1) + apply (erule assms) + apply (erule assms) + apply (erule assms) apply assumption + apply (erule assms) apply assumption +apply(case_tac "x1 = x2") +apply(subgoal_tac + "\q. (q \ {atom x1, atom x2}) \* c \ (supp (([[atom x1, atom x2]]lst. trm1), ([[atom x2]]lst. trm2))) \* q") +apply(erule exE)+ +apply(erule conjE)+ +apply(perm_simp) +apply(rule assms(5)) +apply assumption +apply simp +apply (rule Let2_rename) +apply (simp only: supp_Pair) +apply (simp only: fresh_star_Un_elim) +apply (simp only: supp_Pair) +apply (simp only: fresh_star_Un_elim) +apply(rule at_set_avoiding2) +apply(simp add: finite_supp) +apply(simp add: finite_supp) +apply(simp add: finite_supp) +apply clarify +apply (simp add: fresh_star_def) +apply (simp add: fresh_def supp_Pair supp_Abs) + + apply(subgoal_tac + "\q. (q \ {atom x2}) \* c \ supp (([[atom x2]]lst. trm2), ([[atom x1, atom x2]]lst. trm1), (atom x1)) \* q") + apply(erule exE)+ + apply(erule conjE)+ + apply(rule assms(5)) apply(perm_simp) apply(simp (no_asm) add: fresh_star_insert) apply(rule conjI) -apply(simp add: fresh_star_def) -apply(rotate_tac 3) +apply (simp add: fresh_star_def) +apply(rotate_tac 2) apply(simp add: fresh_star_def) apply(simp) -apply(simp add: foo.eq_iff) -apply(drule supp_perm_eq[symmetric])+ -apply(simp add: atom_eqvt) -apply(rule conjI) -oops - +apply (rule Let2_rename3) +apply (simp add: supp_Pair fresh_star_union) +apply (simp add: supp_Pair fresh_star_union) +apply (simp add: supp_Pair fresh_star_union) +apply clarify +apply (simp add: fresh_star_def supp_atom) +apply(rule at_set_avoiding2) +apply(simp add: finite_supp) +apply(simp add: finite_supp) +apply(simp add: finite_supp) +apply(simp add: fresh_star_def) +apply (simp add: fresh_def supp_Pair supp_Abs supp_atom) +done end diff -r 3696659358c8 -r 6c4372a1f220 Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Wed Nov 24 02:36:21 2010 +0000 +++ b/Nominal/Nominal2_Abs.thy Thu Nov 25 01:18:24 2010 +0000 @@ -553,7 +553,7 @@ lemma [quot_preserve]: assumes q1: "Quotient R1 abs1 rep1" and q2: "Quotient R2 abs2 rep2" - shows "((abs1 ---> id) ---> (abs2 ---> id) ---> prod_fun rep1 rep2 ---> id) prod_fv = prod_fv" + shows "((abs1 ---> id) ---> (abs2 ---> id) ---> map_pair rep1 rep2 ---> id) prod_fv = prod_fv" by (simp add: fun_eq_iff Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]) lemma [mono]: diff -r 3696659358c8 -r 6c4372a1f220 Nominal/nominal_dt_quot.ML --- a/Nominal/nominal_dt_quot.ML Wed Nov 24 02:36:21 2010 +0000 +++ b/Nominal/nominal_dt_quot.ML Thu Nov 25 01:18:24 2010 +0000 @@ -99,7 +99,7 @@ val parser = Scan.repeat (exclude || any) in fun unraw_str s = - s |> explode + s |> raw_explode |> Scan.finite Symbol.stopper parser >> implode |> fst end