diff -r 17b369a73f15 -r cd2aca704279 Nominal/FSet.thy --- a/Nominal/FSet.thy Mon Oct 18 12:15:44 2010 +0100 +++ b/Nominal/FSet.thy Tue Oct 19 10:10:41 2010 +0100 @@ -37,30 +37,30 @@ lists. *} -fun +definition memb :: "'a \ 'a list \ bool" where - "memb x xs \ x \ set xs" + [simp]: "memb x xs \ x \ set xs" -fun +definition sub_list :: "'a list \ 'a list \ bool" where - "sub_list xs ys \ set xs \ set ys" + [simp]: "sub_list xs ys \ set xs \ set ys" -fun +definition card_list :: "'a list \ nat" where - "card_list xs = card (set xs)" + [simp]: "card_list xs = card (set xs)" -fun +definition inter_list :: "'a list \ 'a list \ 'a list" where - "inter_list xs ys = [x \ xs. x \ set xs \ x \ set ys]" + [simp]: "inter_list xs ys = [x \ xs. x \ set xs \ x \ set ys]" -fun +definition diff_list :: "'a list \ 'a list \ 'a list" where - "diff_list xs ys = [x \ xs. x \ set ys]" + [simp]: "diff_list xs ys = [x \ xs. x \ set ys]" definition rsp_fold @@ -1018,7 +1018,7 @@ assume h: "x \ a" then have f: "\ memb x (a # ys)" using d by auto have g: "a # xs \ x # (a # ys)" using e h by auto - show thesis using b f g by (simp del: memb.simps) + show thesis using b f g by (simp del: memb_def) qed qed then show thesis using a c by blast @@ -1158,7 +1158,7 @@ *} no_notation - list_eq (infix "\" 50) -and list_eq2 (infix "\2" 50) + list_eq (infix "\" 50) and + list_eq2 (infix "\2" 50) end