# HG changeset patch # User Christian Urban # Date 1291384267 0 # Node ID 98236fbd8aa6e0da548f7fa43aeff622caf3cd06 # Parent 35c570891a3a107ee913edcaee3dad4a2392d6e6 updated to Isabelle 2nd December diff -r 35c570891a3a -r 98236fbd8aa6 Nominal/Ex/Foo2.thy --- a/Nominal/Ex/Foo2.thy Mon Nov 29 08:01:09 2010 +0000 +++ b/Nominal/Ex/Foo2.thy Fri Dec 03 13:51:07 2010 +0000 @@ -64,7 +64,6 @@ tests by cu *} - lemma yy: assumes a: "p \ bs \ bs = {}" and b: "finite bs" @@ -100,6 +99,8 @@ apply(simp add: mem_permute_iff) done + + lemma Abs_rename_set: fixes x::"'a::fs" assumes a: "(p \ bs) \* (bs, x)" @@ -217,7 +218,7 @@ 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 "\name trm. \{atom name} \* c; y = Lam name trm\ \ P" and "\a1 trm1 a2 trm2. ((set (bn a1)) \ (set (bn a2))) \* c \ y = Let1 a1 trm1 a2 trm2 \ P" and "\x1 x2 trm1 trm2. {atom x1, atom x2} \* c \ y = Let2 x1 x2 trm1 trm2 \ P" shows "P" @@ -230,7 +231,6 @@ apply(assumption) apply(subgoal_tac "\p. (p \ {atom name}) \* (c, [atom name], trm)") apply(erule exE) -apply(rule assms(3)) apply(insert Abs_rename_list)[1] apply(drule_tac x="p" in meta_spec) apply(drule_tac x="[atom name]" in meta_spec) @@ -239,12 +239,14 @@ apply(drule meta_mp) apply(simp) apply(erule exE) -apply(rule exI)+ +apply(rule assms(3)) apply(perm_simp) apply(erule conjE)+ -apply(rule conjI) apply(assumption) -apply(simp add: foo.eq_iff) +apply(clarify) +apply(simp (no_asm) add: foo.eq_iff) +apply(perm_simp) +apply(assumption) apply(rule at_set_avoiding1) apply(simp) apply(simp add: finite_supp) diff -r 35c570891a3a -r 98236fbd8aa6 Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Mon Nov 29 08:01:09 2010 +0000 +++ b/Nominal/Nominal2_Abs.thy Fri Dec 03 13:51:07 2010 +0000 @@ -255,7 +255,7 @@ and 'b abs_res = "(atom set \ 'b::pt)" / "alpha_abs_res" and 'c abs_lst = "(atom list \ 'c::pt)" / "alpha_abs_lst" apply(rule_tac [!] equivpI) - unfolding reflp_def symp_def transp_def + unfolding reflp_def refl_on_def symp_def sym_def transp_def trans_def by (auto intro: alphas_abs_sym alphas_abs_refl alphas_abs_trans simp only:) quotient_definition diff -r 35c570891a3a -r 98236fbd8aa6 Nominal/nominal_dt_alpha.ML --- a/Nominal/nominal_dt_alpha.ML Mon Nov 29 08:01:09 2010 +0000 +++ b/Nominal/nominal_dt_alpha.ML Fri Dec 03 13:51:07 2010 +0000 @@ -593,14 +593,20 @@ (** proves the equivp predicate for all alphas **) +val reflp_def' = + @{lemma "reflp R == !x. R x x" by (simp add: reflp_def refl_on_def)} + +val symp_def' = + @{lemma "symp R == !x y . R x y --> R y x" by (simp add: symp_def sym_def)} + val transp_def' = @{lemma "transp R == !x y. R x y --> (!z. R y z --> R x z)" - by (rule eq_reflection, auto simp add: transp_def)} + by (rule eq_reflection, auto simp add: trans_def transp_def)} fun raw_prove_equivp alphas alpha_bns refl symm trans ctxt = let - val refl' = map (fold_rule @{thms reflp_def[THEN eq_reflection]} o atomize) refl - val symm' = map (fold_rule @{thms symp_def[THEN eq_reflection]} o atomize) symm + val refl' = map (fold_rule [reflp_def'] o atomize) refl + val symm' = map (fold_rule [symp_def'] o atomize) symm val trans' = map (fold_rule [transp_def'] o atomize) trans fun prep_goal t =