diff -r 022a9c26b14f -r 34e0ddfc6dcc src/HOL/Nominal/Examples/Pattern.thy --- a/src/HOL/Nominal/Examples/Pattern.thy Mon Apr 22 10:43:57 2024 +0100 +++ b/src/HOL/Nominal/Examples/Pattern.thy Mon Apr 22 22:08:28 2024 +0100 @@ -355,7 +355,7 @@ by (metis abs_fresh(1) btrm.fresh(2) fresh_star_set supp_fst) qed (auto simp: fresh_star_set) -lemma psubst_nil: "[]\t\ = t" "[]\t'\\<^sub>b = t'" +lemma psubst_nil[simp]: "[]\t\ = t" "[]\t'\\<^sub>b = t'" by (induct t and t' rule: trm_btrm.inducts) (simp_all add: fresh_list_nil) lemma psubst_cons: @@ -367,9 +367,20 @@ lemma psubst_append: "(supp (map fst (\\<^sub>1 @ \\<^sub>2))::name set) \* map snd (\\<^sub>1 @ \\<^sub>2) \ (\\<^sub>1 @ \\<^sub>2)\t\ = \\<^sub>2\\\<^sub>1\t\\" - by (induct \\<^sub>1 arbitrary: t) - (simp_all add: psubst_nil split_paired_all supp_list_cons psubst_cons fresh_star_def - fresh_list_cons fresh_list_append supp_list_append) +proof (induct \\<^sub>1 arbitrary: t) + case Nil + then show ?case + by (auto simp: psubst_nil) +next + case (Cons a \\<^sub>1) + then show ?case + proof (cases a) + case (Pair a b) + with Cons show ?thesis + apply (simp add: supp_list_cons fresh_star_set fresh_list_cons) + by (metis Un_iff fresh_star_set map_append psubst_cons(1) supp_list_append) + qed +qed lemma abs_pat_psubst [simp]: "(supp p::name set) \* \ \ \\\[p]. t\\<^sub>b = (\[p]. \\t\\<^sub>b)"