# HG changeset patch # User Christian Urban # Date 1338462061 -3600 # Node ID ca162f0a7957f93146e2144f05e8c4707ecaaf86 # Parent 7b5db6c237535e8b38e267cc19d4f48bd7fba5a4 added to the simplifier nominal_datatype.fresh lemmas diff -r 7b5db6c23753 -r ca162f0a7957 Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Thu May 31 11:59:56 2012 +0100 +++ b/Nominal/Ex/Lambda.thy Thu May 31 12:01:01 2012 +0100 @@ -299,13 +299,13 @@ lemma forget: shows "atom x \ t \ t[x ::= s] = t" by (nominal_induct t avoiding: x s rule: lam.strong_induct) - (auto simp add: lam.fresh fresh_at_base) + (auto simp add: fresh_at_base) text {* same lemma but with subst.induction *} lemma forget2: shows "atom x \ t \ t[x ::= s] = t" by (induct t x s rule: subst.induct) - (auto simp add: lam.fresh fresh_at_base fresh_Pair) + (auto simp add: fresh_at_base fresh_Pair) lemma fresh_fact: fixes z::"name" @@ -314,7 +314,7 @@ shows "atom z \ t[y ::= s]" using a b by (nominal_induct t avoiding: z y s rule: lam.strong_induct) - (auto simp add: lam.fresh fresh_at_base) + (auto simp add: fresh_at_base) lemma substitution_lemma: assumes a: "x \ y" "atom x \ u" @@ -328,7 +328,7 @@ shows "t[x ::= s] = ((y \ x) \t)[y ::= s]" using a apply (nominal_induct t avoiding: x y s rule: lam.strong_induct) -apply (auto simp add: lam.fresh fresh_at_base) +apply (auto simp add: fresh_at_base) done lemma height_ge_one: @@ -368,7 +368,7 @@ nominal_inductive beta avoids b4: "x" - by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) + by (simp_all add: fresh_star_def fresh_Pair fresh_fact) text {* One-Reduction *} @@ -385,7 +385,7 @@ nominal_inductive One avoids o3: "x" | o4: "x" - by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) + by (simp_all add: fresh_star_def fresh_Pair fresh_fact) lemma One_refl: shows "t \1 t" @@ -610,7 +610,7 @@ lemma var_fresh_subst: "atom x \ s \ atom x \ (t[x ::= s])" - by (induct t x s rule: subst.induct) (auto simp add: lam.supp lam.fresh fresh_at_base) + by (induct t x s rule: subst.induct) (auto simp add: lam.supp fresh_at_base) (* function that evaluates a lambda term *) nominal_primrec @@ -841,7 +841,7 @@ apply(drule_tac x="namea" in meta_spec) apply(drule_tac x="lam" in meta_spec) apply(simp add: fresh_star_Pair) - apply(simp add: fresh_star_def fresh_at_base lam.fresh) + apply(simp add: fresh_star_def fresh_at_base ) apply(auto)[1] apply(simp_all)[44] apply(simp del: Product_Type.prod.inject) @@ -880,7 +880,7 @@ apply(drule_tac x="name" in meta_spec) apply(drule_tac x="lam" in meta_spec) apply(drule_tac x="(name \ namea) \ lama" in meta_spec) - apply(simp add: Abs1_eq_iff fresh_star_def fresh_Pair_elim fresh_at_base lam.fresh flip_def) + apply(simp add: Abs1_eq_iff fresh_star_def fresh_Pair_elim fresh_at_base flip_def) apply (metis Nominal2_Base.swap_commute fresh_permute_iff sort_of_atom_eq swap_atom_simps(2)) apply simp_all apply (simp add: abs_same_binder) diff -r 7b5db6c23753 -r ca162f0a7957 Nominal/Nominal2.thy --- a/Nominal/Nominal2.thy Thu May 31 11:59:56 2012 +0100 +++ b/Nominal/Nominal2.thy Thu May 31 12:01:01 2012 +0100 @@ -522,7 +522,7 @@ ||>> Local_Theory.note ((thms_suffix "supports", []), qsupports_thms) ||>> Local_Theory.note ((thms_suffix "fsupp", []), qfsupp_thms) ||>> Local_Theory.note ((thms_suffix "supp", []), qsupp_constrs) - ||>> Local_Theory.note ((thms_suffix "fresh", []), qfresh_constrs) + ||>> Local_Theory.note ((thms_suffix "fresh", [simp_attr]), qfresh_constrs) ||>> Local_Theory.note ((thms_suffix "perm_bn_simps", []), qperm_bn_simps) ||>> Local_Theory.note ((thms_suffix "bn_finite", []), qbn_finite_thms) ||>> Local_Theory.note ((thms_suffix "perm_bn_alpha", []), qperm_bn_alpha_thms)