major reorganisation of fset (renamed fset_to_set to fset, changed the definition of list_eq and fcard_raw)
authorChristian Urban <urbanc@in.tum.de>
Thu, 14 Oct 2010 04:14:22 +0100
changeset 2524 693562f03eee
parent 2523 e903c32ec24f
child 2525 c848f93807b9
major reorganisation of fset (renamed fset_to_set to fset, changed the definition of list_eq and fcard_raw)
IsaMakefile
Nominal/Equivp.thy
Nominal/Ex/SingleLet.thy
Nominal/Ex/TypeSchemes.thy
Nominal/FSet.thy
Nominal/Nominal2_FSet.thy
Nominal/nominal_dt_rawfuns.ML
--- a/IsaMakefile	Wed Oct 13 22:55:58 2010 +0100
+++ b/IsaMakefile	Thu Oct 14 04:14:22 2010 +0100
@@ -1,10 +1,10 @@
 
 ## targets
 
-default: test
+default: tests
 images: 
 
-all: tests paper pearl pearl-jv qpaper
+all: tests paper pearl pearl-jv qpaper slides
 
 
 ## global settings
@@ -107,7 +107,7 @@
 	cd Slides/generated4 ; $(ISABELLE_TOOL) latex -o pdf root.beamer.tex
 	cp Slides/generated4/root.beamer.pdf Slides/slides4.pdf 
 
-slides: slides4
+slides: slides1 slides2 slides3 slides4
 
 
 
--- a/Nominal/Equivp.thy	Wed Oct 13 22:55:58 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,185 +0,0 @@
-theory Equivp
-imports "Abs" "Perm" "Tacs" "Rsp"
-begin
-
-lemma not_in_union: "c \<notin> a \<union> b \<equiv> (c \<notin> a \<and> c \<notin> b)"
-by auto
-
-ML {*
-fun supports_tac perm =
-  simp_tac (HOL_ss addsimps @{thms supports_def not_in_union} @ perm) THEN_ALL_NEW (
-    REPEAT o rtac allI THEN' REPEAT o rtac impI THEN' split_conj_tac THEN'
-    asm_full_simp_tac (HOL_ss addsimps @{thms fresh_def[symmetric]
-      swap_fresh_fresh fresh_atom swap_at_base_simps(3) swap_atom_image_fresh
-      supp_fset_to_set supp_fmap_atom}))
-*}
-
-ML {*
-fun mk_supp ty x =
-  Const (@{const_name supp}, ty --> @{typ "atom set"}) $ x
-*}
-
-ML {*
-fun mk_supports_eq thy cnstr =
-let
-  val (tys, ty) = (strip_type o fastype_of) cnstr
-  val names = Datatype_Prop.make_tnames tys
-  val frees = map Free (names ~~ tys)
-  val rhs = list_comb (cnstr, frees)
-
-  fun mk_supp_arg (x, ty) =
-    if is_atom thy ty then mk_supp @{typ atom} (mk_atom_ty ty x) else
-    if is_atom_set thy ty then mk_supp @{typ "atom set"} (mk_atom_set x) else
-    if is_atom_fset thy ty then mk_supp @{typ "atom set"} (mk_atom_fset x)
-    else mk_supp ty x
-  val lhss = map mk_supp_arg (frees ~~ tys)
-  val supports = Const(@{const_name "supports"}, @{typ "atom set"} --> ty --> @{typ bool})
-  val eq = HOLogic.mk_Trueprop (supports $ fold_union lhss $ rhs)
-in
-  (names, eq)
-end
-*}
-
-ML {*
-fun prove_supports ctxt perms cnst =
-let
-  val (names, eq) = mk_supports_eq ctxt cnst
-in
-  Goal.prove ctxt names [] eq (fn _ => supports_tac perms 1)
-end
-*}
-
-ML {*
-fun mk_fs tys =
-let
-  val names = Datatype_Prop.make_tnames tys
-  val frees = map Free (names ~~ tys)
-  val supps = map2 mk_supp tys frees
-  val fin_supps = map (fn x => @{term "finite :: atom set \<Rightarrow> bool"} $ x) supps
-in
-  (names, HOLogic.mk_Trueprop (mk_conjl fin_supps))
-end
-*}
-
-ML {*
-fun fs_tac induct supports = rtac induct THEN_ALL_NEW (
-  rtac @{thm supports_finite} THEN' resolve_tac supports) THEN_ALL_NEW
-  asm_full_simp_tac (HOL_ss addsimps @{thms supp_atom supp_atom_image supp_fset_to_set
-    supp_fmap_atom finite_insert finite.emptyI finite_Un finite_supp})
-*}
-
-ML {*
-fun prove_fs ctxt induct supports tys =
-let
-  val (names, eq) = mk_fs tys
-in
-  Goal.prove ctxt names [] eq (fn _ => fs_tac induct supports 1)
-end
-*}
-
-ML {*
-fun mk_supp x = Const (@{const_name supp}, fastype_of x --> @{typ "atom set"}) $ x;
-
-fun mk_supp_neq arg (fv, alpha) =
-let
-  val collect = Const ("Collect", @{typ "(atom \<Rightarrow> bool) \<Rightarrow> atom \<Rightarrow> bool"});
-  val ty = fastype_of arg;
-  val perm = Const ("Nominal2_Base.pt_class.permute", @{typ perm} --> ty --> ty);
-  val finite = @{term "finite :: atom set \<Rightarrow> bool"}
-  val rhs = collect $ Abs ("a", @{typ atom},
-    HOLogic.mk_not (finite $
-      (collect $ Abs ("b", @{typ atom},
-        HOLogic.mk_not (alpha $ (perm $ (@{term swap} $ Bound 1 $ Bound 0) $ arg) $ arg)))))
-in
-  HOLogic.mk_eq (fv $ arg, rhs)
-end;
-
-fun supp_eq fv_alphas_lst =
-let
-  val (fvs_alphas, ls) = split_list fv_alphas_lst;
-  val (fv_ts, _) = split_list fvs_alphas;
-  val tys = map (domain_type o fastype_of) fv_ts;
-  val names = Datatype_Prop.make_tnames tys;
-  val args = map Free (names ~~ tys);
-  fun supp_eq_arg ((fv, arg), l) =
-    mk_conjl
-      ((HOLogic.mk_eq (fv $ arg, mk_supp arg)) ::
-       (map (mk_supp_neq arg) l))
-  val eqs = mk_conjl (map supp_eq_arg ((fv_ts ~~ args) ~~ ls))
-in
-  (names, HOLogic.mk_Trueprop eqs)
-end
-*}
-
-ML {*
-fun combine_fv_alpha_bns (fv_ts_nobn, fv_ts_bn) (alpha_ts_nobn, alpha_ts_bn) bn_nos =
-if length fv_ts_bn < length alpha_ts_bn then
-  (fv_ts_nobn ~~ alpha_ts_nobn) ~~ (replicate (length fv_ts_nobn) [])
-else let
-  val fv_alpha_nos = 0 upto (length fv_ts_nobn - 1);
-  fun filter_fn i (x, j) = if j = i then SOME x else NONE;
-  val fv_alpha_bn_nos = (fv_ts_bn ~~ alpha_ts_bn) ~~ bn_nos;
-  val fv_alpha_bn_all = map (fn i => map_filter (filter_fn i) fv_alpha_bn_nos) fv_alpha_nos;
-in
-  (fv_ts_nobn ~~ alpha_ts_nobn) ~~ fv_alpha_bn_all
-end
-*}
-
-(* TODO: this is a hack, it assumes that only one type of Abs's is present
-   in the type and chooses this supp_abs. Additionally single atoms are
-   treated properly. *)
-ML {*
-fun choose_alpha_abs eqiff =
-let
-  fun exists_subterms f ts = member (op =) (map (exists_subterm f) ts) true;
-  val terms = map prop_of eqiff;
-  fun check cname = exists_subterms (fn x => fst(dest_Const x) = cname handle _ => false) terms
-  val no =
-    if check @{const_name alpha_lst} then 2 else
-    if check @{const_name alpha_res} then 1 else
-    if check @{const_name alpha_gen} then 0 else
-    error "Failure choosing supp_abs"
-in
-  nth @{thms supp_abs[symmetric]} no
-end
-*}
-lemma supp_abs_atom: "supp (Abs {atom a} (x :: 'a :: fs)) = supp x - {atom a}"
-by (rule supp_abs(1))
-
-lemma supp_abs_sum:
-  "supp (Abs x (a :: 'a :: fs)) \<union> supp (Abs x (b :: 'b :: fs)) = supp (Abs x (a, b))"
-  "supp (Abs_res x (a :: 'a :: fs)) \<union> supp (Abs_res x (b :: 'b :: fs)) = supp (Abs_res x (a, b))"
-  "supp (Abs_lst y (a :: 'a :: fs)) \<union> supp (Abs_lst y (b :: 'b :: fs)) = supp (Abs_lst y (a, b))"
-  apply (simp_all add: supp_abs supp_Pair)
-  apply blast+
-  done
-
-
-ML {*
-fun supp_eq_tac ind fv perm eqiff ctxt =
-  rtac ind THEN_ALL_NEW
-  asm_full_simp_tac (HOL_basic_ss addsimps fv) THEN_ALL_NEW
-  asm_full_simp_tac (HOL_basic_ss addsimps @{thms supp_abs_atom[symmetric]}) THEN_ALL_NEW
-  asm_full_simp_tac (HOL_basic_ss addsimps [choose_alpha_abs eqiff]) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms supp_abs_sum}) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms supp_def}) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps (@{thms permute_abs} @ perm)) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps (@{thms Abs_eq_iff} @ eqiff)) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms alphas3 alphas2}) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms alphas}) THEN_ALL_NEW
-  asm_full_simp_tac (HOL_basic_ss addsimps (@{thm supp_Pair} :: sym_eqvts ctxt)) THEN_ALL_NEW
-  asm_full_simp_tac (HOL_basic_ss addsimps (@{thm Pair_eq} :: all_eqvts ctxt)) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms supp_at_base[symmetric,simplified supp_def]}) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms Collect_disj_eq[symmetric]}) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms infinite_Un[symmetric]}) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms Collect_disj_eq[symmetric]}) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms de_Morgan_conj[symmetric]}) THEN_ALL_NEW
-  simp_tac (HOL_basic_ss addsimps @{thms ex_simps(1,2)[symmetric]}) THEN_ALL_NEW
-  simp_tac (HOL_ss addsimps @{thms Collect_const finite.emptyI})
-*}
-
-
-
-
-
-end
--- a/Nominal/Ex/SingleLet.thy	Wed Oct 13 22:55:58 2010 +0100
+++ b/Nominal/Ex/SingleLet.thy	Thu Oct 14 04:14:22 2010 +0100
@@ -157,10 +157,7 @@
   apply(simp add: finite_supp)
   apply(simp add: finite_supp)
   oops
-  (*apply(simp add: Abs_fresh_star)*)
-sorry
-  
-done
+
 
 end
 
--- a/Nominal/Ex/TypeSchemes.thy	Wed Oct 13 22:55:58 2010 +0100
+++ b/Nominal/Ex/TypeSchemes.thy	Thu Oct 14 04:14:22 2010 +0100
@@ -56,7 +56,7 @@
 lemma strong_induct:
   assumes a1: "\<And>name b. P b (Var name)"
   and     a2: "\<And>t1 t2 b. \<lbrakk>\<And>c. P c t1; \<And>c. P c t2\<rbrakk> \<Longrightarrow> P b (Fun t1 t2)"
-  and     a3: "\<And>fset t b. \<lbrakk>\<And>c. P c t; fset_to_set (fmap atom fset) \<sharp>* b\<rbrakk> \<Longrightarrow> P' b (All fset t)"
+  and     a3: "\<And>fset t b. \<lbrakk>\<And>c. P c t; fset (fmap atom fset) \<sharp>* b\<rbrakk> \<Longrightarrow> P' b (All fset t)"
   shows "P (a :: 'a :: pt) t \<and> P' (d :: 'b :: {fs}) ts "
 proof -
   have " (\<forall>p a. P a (p \<bullet> t)) \<and> (\<forall>p d. P' d (p \<bullet> ts))"
@@ -69,7 +69,7 @@
     apply simp
     apply (rule allI)
     apply (rule allI)
-    apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (fset_to_set (fmap atom (p \<bullet> fset)))) \<sharp>* d \<and> supp (p \<bullet> All fset ty) \<sharp>* pa)")
+    apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (fset (fmap atom (p \<bullet> fset)))) \<sharp>* d \<and> supp (p \<bullet> All fset ty) \<sharp>* pa)")
     apply clarify
     apply(rule_tac t="p \<bullet> All fset ty" and 
                    s="pa \<bullet> (p \<bullet> All fset ty)" in subst)
@@ -81,7 +81,7 @@
     apply simp
     apply (simp add: eqvts eqvts_raw)
     apply (rule at_set_avoiding2)
-    apply (simp add: fin_fset_to_set)
+    apply (simp add: fin_fset)
     apply (simp add: finite_supp)
     apply (simp add: eqvts finite_supp)
     apply (rule_tac p=" -p" in permute_boolE)
@@ -142,7 +142,7 @@
 assumes
     s1: "subst \<theta> (Var n) = lookup \<theta> n"
 and s2: "subst \<theta> (Fun l r) = Fun (subst \<theta> l) (subst \<theta> r)"
-and s3: "fset_to_set (fmap atom xs) \<sharp>* \<theta> \<Longrightarrow> substs \<theta> (All xs t) = All xs (subst \<theta> t)"
+and s3: "fset (fmap atom xs) \<sharp>* \<theta> \<Longrightarrow> substs \<theta> (All xs t) = All xs (subst \<theta> t)"
 begin
 
 lemma subst_ty:
@@ -163,7 +163,7 @@
   apply (simp_all add: fresh_star_prod_elim)
   apply (drule fresh_star_atom)
   apply (simp add: fresh_def ty_tys.fv[simplified ty_tys.supp])
-  apply (subgoal_tac "atom a \<notin> fset_to_set (fmap atom fset)")
+  apply (subgoal_tac "atom a \<notin> fset (fmap atom fset)")
   apply blast
   apply (metis supp_finite_atom_set finite_fset)
   done
@@ -197,7 +197,7 @@
   apply (fold fresh_def)[1]
   apply (rule mp[OF subst_lemma_pre])
   apply (simp add: fresh_Pair)
-  apply (subgoal_tac "atom a \<notin> (fset_to_set (fmap atom fset))")
+  apply (subgoal_tac "atom a \<notin> (fset (fmap atom fset))")
   apply blast
   apply (metis supp_finite_atom_set finite_fset)
   done
--- a/Nominal/FSet.thy	Wed Oct 13 22:55:58 2010 +0100
+++ b/Nominal/FSet.thy	Thu Oct 14 04:14:22 2010 +0100
@@ -14,7 +14,7 @@
 fun
   list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
 where
-  "list_eq xs ys = (\<forall>x. x \<in> set xs \<longleftrightarrow> x \<in> set ys)"
+  "list_eq xs ys = (set xs = set ys)"
 
 lemma list_eq_equivp:
   shows "equivp list_eq"
@@ -38,34 +38,25 @@
 definition
   sub_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
 where
-  "sub_list xs ys \<equiv> (\<forall>x. x \<in> set xs \<longrightarrow> x \<in> set ys)"
+  "sub_list xs ys \<equiv> set xs \<subseteq> set ys"
 
-fun
+definition
   fcard_raw :: "'a list \<Rightarrow> nat"
 where
-  fcard_raw_nil:  "fcard_raw [] = 0"
-| fcard_raw_cons: "fcard_raw (x # xs) = 
-    (if memb x xs then fcard_raw xs else Suc (fcard_raw xs))"
+  "fcard_raw xs = card (set xs)"
 
 primrec
   finter_raw :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 where
-  "finter_raw [] l = []"
-| "finter_raw (h # t) l =
-    (if memb h l then h # (finter_raw t l) else finter_raw t l)"
-
-primrec
-  delete_raw :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
-where
-  "delete_raw [] x = []"
-| "delete_raw (a # xs) x = 
-    (if (a = x) then delete_raw xs x else a # (delete_raw xs x))"
+  "finter_raw [] ys = []"
+| "finter_raw (x # xs) ys =
+    (if x \<in> set ys then x # (finter_raw xs ys) else finter_raw xs ys)"
 
 primrec
   fminus_raw :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 where
-  "fminus_raw l [] = l"
-| "fminus_raw l (h # t) = fminus_raw (delete_raw l h) t"
+  "fminus_raw ys [] = ys"
+| "fminus_raw ys (x # xs) = fminus_raw (removeAll x ys) xs"
 
 definition
   rsp_fold
@@ -78,7 +69,7 @@
   "ffold_raw f z [] = z"
 | "ffold_raw f z (a # xs) =
      (if (rsp_fold f) then
-       if memb a xs then ffold_raw f z xs
+       if a \<in> set xs then ffold_raw f z xs
        else f a (ffold_raw f z xs)
      else z)"
 
@@ -100,12 +91,14 @@
   shows "Quotient (list_all2 op \<approx>) (map abs_fset) (map rep_fset)"
   by (fact list_quotient[OF Quotient_fset])
 
+(*
 lemma set_in_eq: "(\<forall>e. ((e \<in> xs) \<longleftrightarrow> (e \<in> ys))) \<equiv> xs = ys"
   by (rule eq_reflection) auto
+*)
 
 lemma map_rel_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba"
   unfolding list_eq.simps
-  by (simp only: set_map set_in_eq)
+  by (simp only: set_map)
 
 lemma quotient_compose_list[quot_thm]:
   shows  "Quotient ((list_all2 op \<approx>) OOO (op \<approx>))
@@ -165,13 +158,14 @@
 text {* Respectfullness *}
 
 lemma append_rsp[quot_respect]:
-  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
-  apply(simp del: list_eq.simps)
-  by auto 
+  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) append append"
+  by (simp)
 
+(*
 lemma append_rsp_unfolded:
   "\<lbrakk>x \<approx> y; v \<approx> w\<rbrakk> \<Longrightarrow> x @ v \<approx> y @ w"
   by auto
+*)
 
 lemma [quot_respect]:
   shows "(op \<approx> ===> op \<approx> ===> op =) sub_list sub_list"
@@ -186,7 +180,7 @@
   by simp
 
 lemma cons_rsp[quot_respect]:
-  shows "(op = ===> op \<approx> ===> op \<approx>) op # op #"
+  shows "(op = ===> op \<approx> ===> op \<approx>) Cons Cons"
   by simp
 
 lemma map_rsp[quot_respect]:
@@ -209,84 +203,35 @@
   shows "memb x (y # xs) = (x = y \<or> memb x xs)"
   by (induct xs) (auto simp add: memb_def)
 
-lemma memb_finter_raw:
-  "memb x (finter_raw xs ys) \<longleftrightarrow> memb x xs \<and> memb x ys"
-  by (induct xs) (auto simp add: not_memb_nil memb_cons_iff)
-
-lemma [quot_respect]:
-  "(op \<approx> ===> op \<approx> ===> op \<approx>) finter_raw finter_raw"
-  by (simp add: memb_def[symmetric] memb_finter_raw)
-
-lemma memb_delete_raw:
-  "memb x (delete_raw xs y) = (memb x xs \<and> x \<noteq> y)"
-  by (induct xs arbitrary: x y) (auto simp add: memb_def)
-
-lemma [quot_respect]:
-  "(op \<approx> ===> op = ===> op \<approx>) delete_raw delete_raw"
-  by (simp add: memb_def[symmetric] memb_delete_raw)
-
-lemma fminus_raw_memb: "memb x (fminus_raw xs ys) = (memb x xs \<and> \<not> memb x ys)"
-  by (induct ys arbitrary: xs)
-     (simp_all add: not_memb_nil memb_delete_raw memb_cons_iff)
+lemma set_finter_raw[simp]:
+  "set (finter_raw xs ys) = set xs \<inter> set ys"
+  by (induct xs) (auto simp add: memb_def)
 
 lemma [quot_respect]:
-  "(op \<approx> ===> op \<approx> ===> op \<approx>) fminus_raw fminus_raw"
-  by (simp add: memb_def[symmetric] fminus_raw_memb)
+  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) finter_raw finter_raw"
+  by (simp)
 
-lemma fcard_raw_gt_0:
-  assumes a: "x \<in> set xs"
-  shows "0 < fcard_raw xs"
-  using a by (induct xs) (auto simp add: memb_def)
-
-lemma fcard_raw_delete_one:
-  shows "fcard_raw ([x \<leftarrow> xs. x \<noteq> y]) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)"
-  by (induct xs) (auto dest: fcard_raw_gt_0 simp add: memb_def)
+(*
+lemma memb_removeAll:
+  "memb x (removeAll y xs) \<longleftrightarrow> memb x xs \<and> x \<noteq> y"
+  by (induct xs arbitrary: x y) (auto simp add: memb_def)
+*)
 
-lemma fcard_raw_rsp_aux:
-  assumes a: "xs \<approx> ys"
-  shows "fcard_raw xs = fcard_raw ys"
-  using a
-  proof (induct xs arbitrary: ys)
-    case Nil
-    show ?case using Nil.prems by simp
-  next
-    case (Cons a xs)
-    have a: "a # xs \<approx> ys" by fact
-    have b: "\<And>ys. xs \<approx> ys \<Longrightarrow> fcard_raw xs = fcard_raw ys" by fact
-    show ?case proof (cases "a \<in> set xs")
-      assume c: "a \<in> set xs"
-      have "\<forall>x. (x \<in> set xs) = (x \<in> set ys)"
-      proof (intro allI iffI)
-        fix x
-        assume "x \<in> set xs"
-        then show "x \<in> set ys" using a by auto
-      next
-        fix x
-        assume d: "x \<in> set ys"
-        have e: "(x \<in> set (a # xs)) = (x \<in> set ys)" using a by simp
-        show "x \<in> set xs" using c d e unfolding list_eq.simps by simp blast
-      qed
-      then show ?thesis using b c by (simp add: memb_def)
-    next
-      assume c: "a \<notin> set xs"
-      have d: "xs \<approx> [x\<leftarrow>ys . x \<noteq> a] \<Longrightarrow> fcard_raw xs = fcard_raw [x\<leftarrow>ys . x \<noteq> a]" using b by simp
-      have "Suc (fcard_raw xs) = fcard_raw ys"
-      proof (cases "a \<in> set ys")
-        assume e: "a \<in> set ys"
-        have f: "\<forall>x. (x \<in> set xs) = (x \<in> set ys \<and> x \<noteq> a)" using a c
-          by (auto simp add: fcard_raw_delete_one)
-        have "fcard_raw ys = Suc (fcard_raw ys - 1)" by (rule Suc_pred'[OF fcard_raw_gt_0]) (rule e)
-        then show ?thesis using d e f by (simp_all add: fcard_raw_delete_one memb_def)
-      next
-        case False then show ?thesis using a c d by auto
-      qed
-      then show ?thesis using a c d by (simp add: memb_def)
-  qed
-qed
+lemma [quot_respect]:
+  shows "(op = ===> op \<approx> ===> op \<approx>) removeAll removeAll"
+  by (simp)
+
+lemma set_fminus_raw[simp]: 
+  "set (fminus_raw xs ys) = (set xs - set ys)"
+  by (induct ys arbitrary: xs) (auto)
+
+lemma [quot_respect]:
+  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) fminus_raw fminus_raw"
+  by simp
 
 lemma fcard_raw_rsp[quot_respect]:
   shows "(op \<approx> ===> op =) fcard_raw fcard_raw"
-  by (simp add: fcard_raw_rsp_aux)
+  by (simp add: fcard_raw_def)
 
 lemma memb_absorb:
   shows "memb x xs \<Longrightarrow> x # xs \<approx> xs"
@@ -296,53 +241,39 @@
   "(\<forall>x. \<not> memb x xs) = (xs \<approx> [])"
   by (simp add: memb_def)
 
-lemma not_memb_delete_raw_ident:
-  shows "\<not> memb x xs \<Longrightarrow> delete_raw xs x = xs"
+lemma not_memb_removeAll_ident:
+  shows "\<not> memb x xs \<Longrightarrow> removeAll x xs = xs"
   by (induct xs) (auto simp add: memb_def)
 
 lemma memb_commute_ffold_raw:
-  "rsp_fold f \<Longrightarrow> memb h b \<Longrightarrow> ffold_raw f z b = f h (ffold_raw f z (delete_raw b h))"
+  "rsp_fold f \<Longrightarrow> h \<in> set b \<Longrightarrow> ffold_raw f z b = f h (ffold_raw f z (removeAll h b))"
   apply (induct b)
-  apply (simp_all add: not_memb_nil)
-  apply (auto)
-  apply (simp_all add: memb_delete_raw not_memb_delete_raw_ident rsp_fold_def  memb_cons_iff)
+  apply (auto simp add: rsp_fold_def)
   done
 
 lemma ffold_raw_rsp_pre:
-  "\<forall>e. memb e a = memb e b \<Longrightarrow> ffold_raw f z a = ffold_raw f z b"
+  "set a = set b \<Longrightarrow> ffold_raw f z a = ffold_raw f z b"
   apply (induct a arbitrary: b)
-  apply (simp add: memb_absorb memb_def none_memb_nil)
   apply (simp)
+  apply (simp (no_asm_use))
   apply (rule conjI)
   apply (rule_tac [!] impI)
   apply (rule_tac [!] conjI)
   apply (rule_tac [!] impI)
-  apply (subgoal_tac "\<forall>e. memb e a2 = memb e b")
-  apply (simp)
-  apply (simp add: memb_cons_iff memb_def)
-  apply (auto)[1]
-  apply (drule_tac x="e" in spec)
-  apply (blast)
-  apply (case_tac b)
-  apply (simp_all)
-  apply (subgoal_tac "ffold_raw f z b = f a1 (ffold_raw f z (delete_raw b a1))")
-  apply (simp only:)
-  apply (rule_tac f="f a1" in arg_cong)
-  apply (subgoal_tac "\<forall>e. memb e a2 = memb e (delete_raw b a1)")
-  apply (simp)
-  apply (simp add: memb_delete_raw)
-  apply (auto simp add: memb_cons_iff)[1]
-  apply (erule memb_commute_ffold_raw)
-  apply (drule_tac x="a1" in spec)
-  apply (simp add: memb_cons_iff)
-  apply (simp add: memb_cons_iff)
-  apply (case_tac b)
-  apply (simp_all)
-  done
+  apply (metis insert_absorb)
+  apply (metis List.insert_def List.set.simps(2) List.set_insert ffold_raw.simps(2))
+  apply (metis Diff_insert_absorb insertI1 memb_commute_ffold_raw set_removeAll)
+  apply(drule_tac x="removeAll a1 b" in meta_spec)
+  apply(auto)
+  apply(drule meta_mp)
+  apply(blast)
+  by (metis List.set.simps(2) emptyE ffold_raw.simps(2) in_listsp_conv_set listsp.simps mem_def)
 
 lemma [quot_respect]:
-  "(op = ===> op = ===> op \<approx> ===> op =) ffold_raw ffold_raw"
-  by (simp add: memb_def[symmetric] ffold_raw_rsp_pre)
+  shows "(op = ===> op = ===> op \<approx> ===> op =) ffold_raw ffold_raw"
+  unfolding fun_rel_def
+  apply(auto intro: ffold_raw_rsp_pre)
+  done
 
 lemma concat_rsp_pre:
   assumes a: "list_all2 op \<approx> x x'"
@@ -366,9 +297,11 @@
   assume a: "list_all2 op \<approx> a ba"
   assume b: "ba \<approx> bb"
   assume c: "list_all2 op \<approx> bb b"
-  have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
+  have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" 
+  proof
     fix x
-    show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
+    show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" 
+    proof
       assume d: "\<exists>xa\<in>set a. x \<in> set xa"
       show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d])
     next
@@ -379,11 +312,10 @@
       show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
     qed
   qed
-  then show "concat a \<approx> concat b" by simp
+  then show "concat a \<approx> concat b" by auto
 qed
 
 
-
 lemma concat_rsp_unfolded:
   "\<lbrakk>list_all2 op \<approx> a ba; ba \<approx> bb; list_all2 op \<approx> bb b\<rbrakk> \<Longrightarrow> concat a \<approx> concat b"
 proof -
@@ -404,11 +336,11 @@
       show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
     qed
   qed
-  then show "concat a \<approx> concat b" by simp
+  then show "concat a \<approx> concat b" by auto
 qed
 
 lemma [quot_respect]:
-  "((op =) ===> op \<approx> ===> op \<approx>) filter filter"
+  shows "((op =) ===> op \<approx> ===> op \<approx>) filter filter"
   by auto
 
 text {* Distributive lattice with bot *}
@@ -439,11 +371,11 @@
 
 lemma sub_list_inter_left:
   shows "sub_list (finter_raw x y) x"
-  by (simp add: sub_list_def memb_def[symmetric] memb_finter_raw)
+  by (simp add: sub_list_def)
 
 lemma sub_list_inter_right:
   shows "sub_list (finter_raw x y) y"
-  by (simp add: sub_list_def memb_def[symmetric] memb_finter_raw)
+  by (simp add: sub_list_def)
 
 lemma sub_list_list_eq:
   "sub_list x y \<Longrightarrow> sub_list y x \<Longrightarrow> list_eq x y"
@@ -455,14 +387,12 @@
 
 lemma sub_list_inter:
   "sub_list x y \<Longrightarrow> sub_list x z \<Longrightarrow> sub_list x (finter_raw y z)"
-  by (simp add: sub_list_def memb_def[symmetric] memb_finter_raw)
+  by (simp add: sub_list_def)
 
 lemma append_inter_distrib:
   "x @ (finter_raw y z) \<approx> finter_raw (x @ y) (x @ z)"
   apply (induct x)
-  apply (simp_all add: memb_def)
-  apply (simp add: memb_def[symmetric] memb_finter_raw)
-  apply (auto simp add: memb_def)
+  apply (auto)
   done
 
 instantiation fset :: (type) "{bounded_lattice_bot,distrib_lattice,minus}"
@@ -487,44 +417,45 @@
   "xs |\<subseteq>| ys \<equiv> xs \<le> ys"
 
 definition
-  less_fset:
-  "(xs :: 'a fset) < ys \<equiv> xs \<le> ys \<and> xs \<noteq> ys"
+  less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool"
+where  
+  "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)"
 
 abbreviation
-  f_subset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50)
+  fsubset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50)
 where
   "xs |\<subset>| ys \<equiv> xs < ys"
 
 quotient_definition
-  "sup  \<Colon> ('a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset)"
+  "sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
 is
-  "(op @) \<Colon> ('a list \<Rightarrow> 'a list \<Rightarrow> 'a list)"
+  "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 
 abbreviation
   funion (infixl "|\<union>|" 65)
 where
-  "xs |\<union>| ys \<equiv> sup (xs :: 'a fset) ys"
+  "xs |\<union>| ys \<equiv> sup xs (ys::'a fset)"
 
 quotient_definition
-  "inf \<Colon> ('a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset)"
+  "inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
 is
-  "finter_raw \<Colon> ('a list \<Rightarrow> 'a list \<Rightarrow> 'a list)"
+  "finter_raw :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 
 abbreviation
   finter (infixl "|\<inter>|" 65)
 where
-  "xs |\<inter>| ys \<equiv> inf (xs :: 'a fset) ys"
+  "xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)"
 
 quotient_definition
-  "minus \<Colon> ('a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset)"
+  "minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
 is
-  "fminus_raw \<Colon> ('a list \<Rightarrow> 'a list \<Rightarrow> 'a list)"
+  "fminus_raw :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 
 instance
 proof
   fix x y z :: "'a fset"
   show "(x |\<subset>| y) = (x |\<subseteq>| y \<and> \<not> y |\<subseteq>| x)"
-    unfolding less_fset by (lifting sub_list_not_eq)
+    unfolding less_fset_def by (lifting sub_list_not_eq)
   show "x |\<subseteq>| x" by (lifting sub_list_refl)
   show "{||} |\<subseteq>| x" by (lifting sub_list_empty)
   show "x |\<subseteq>| x |\<union>| y" by (lifting sub_list_append_left)
@@ -560,7 +491,7 @@
 
 quotient_definition
   "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is "op #"
+is "Cons"
 
 syntax
   "@Finset"     :: "args => 'a fset"  ("{|(_)|}")
@@ -583,20 +514,18 @@
 
 quotient_definition
   "fcard :: 'a fset \<Rightarrow> nat"
-is
-  "fcard_raw"
+  is fcard_raw
 
 quotient_definition
   "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
-is
- "map"
+  is map
 
 quotient_definition
-  "fdelete :: 'a fset \<Rightarrow> 'a \<Rightarrow> 'a fset"
-  is "delete_raw"
+  "fdelete :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+  is removeAll
 
 quotient_definition
-  "fset_to_set :: 'a fset \<Rightarrow> 'a set"
+  "fset :: 'a fset \<Rightarrow> 'a set"
   is "set"
 
 quotient_definition
@@ -622,9 +551,8 @@
   by simp
 
 lemma [quot_respect]:
-  "(op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) op # op #"
+  shows "(op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) Cons Cons"
   apply auto
-  apply (simp add: set_in_eq)
   apply (rule_tac b="x # b" in pred_compI)
   apply auto
   apply (rule_tac b="x # ba" in pred_compI)
@@ -723,53 +651,25 @@
 
 lemma singleton_list_eq:
   shows "[x] \<approx> [y] \<longleftrightarrow> x = y"
-  by (simp add:) auto
+  by (simp)
 
 lemma sub_list_cons:
   "sub_list (x # xs) ys = (memb x ys \<and> sub_list xs ys)"
   by (auto simp add: memb_def sub_list_def)
 
-lemma fminus_raw_red: "fminus_raw (x # xs) ys = (if memb x ys then fminus_raw xs ys else x # (fminus_raw xs ys))"
-  by (induct ys arbitrary: xs x)
-     (simp_all add: not_memb_nil memb_delete_raw memb_cons_iff)
+lemma fminus_raw_red: 
+  "fminus_raw (x # xs) ys = (if x \<in> set ys then fminus_raw xs ys else x # (fminus_raw xs ys))"
+  by (induct ys arbitrary: xs x) (simp_all)
 
 text {* Cardinality of finite sets *}
 
+(* used in memb_card_not_0 *)
 lemma fcard_raw_0:
   shows "fcard_raw xs = 0 \<longleftrightarrow> xs \<approx> []"
-  by (induct xs) (auto simp add: memb_def)
-
-lemma fcard_raw_not_memb:
-  shows "\<not> memb x xs \<longleftrightarrow> fcard_raw (x # xs) = Suc (fcard_raw xs)"
-  by auto
-
-lemma fcard_raw_suc:
-  assumes a: "fcard_raw xs = Suc n"
-  shows "\<exists>x ys. \<not> (memb x ys) \<and> xs \<approx> (x # ys) \<and> fcard_raw ys = n"
-  using a
-  by (induct xs) (auto simp add: memb_def split_ifs)
-
-lemma singleton_fcard_1:
-  shows "set xs = {x} \<Longrightarrow> fcard_raw xs = 1"
-  by (induct xs) (auto simp add: memb_def subset_insert)
+  unfolding fcard_raw_def
+  by (induct xs) (auto)
 
-lemma fcard_raw_1:
-  shows "fcard_raw xs = 1 \<longleftrightarrow> (\<exists>x. xs \<approx> [x])"
-  apply (auto dest!: fcard_raw_suc)
-  apply (simp add: fcard_raw_0)
-  apply (rule_tac x="x" in exI)
-  apply simp
-  apply (subgoal_tac "set xs = {x}")
-  apply (drule singleton_fcard_1)
-  apply auto
-  done
-
-lemma fcard_raw_suc_memb:
-  assumes a: "fcard_raw A = Suc n"
-  shows "\<exists>a. memb a A"
-  using a
-  by (induct A) (auto simp add: memb_def)
-
+(* used in list_eq2_equiv *)
 lemma memb_card_not_0:
   assumes a: "memb a A"
   shows "\<not>(fcard_raw A = 0)"
@@ -779,7 +679,11 @@
   then show ?thesis using fcard_raw_0[of A] by simp
 qed
 
-text {* fmap *}
+
+
+section {* fmap *}
+
+(* there is another fmap section below *)
 
 lemma map_append:
   "map f (xs @ ys) \<approx> (map f xs) @ (map f ys)"
@@ -830,32 +734,29 @@
   then show thesis using a c by blast
 qed
 
+
 section {* deletion *}
 
-lemma memb_delete_raw_ident:
-  shows "\<not> memb x (delete_raw xs x)"
+lemma memb_removeAll_ident:
+  shows "\<not> memb x (removeAll x xs)"
   by (induct xs) (auto simp add: memb_def)
 
-lemma fset_raw_delete_raw_cases:
-  "xs = [] \<or> (\<exists>x. memb x xs \<and> xs \<approx> x # delete_raw xs x)"
+lemma fset_raw_removeAll_cases:
+  "xs = [] \<or> (\<exists>x. memb x xs \<and> xs \<approx> x # removeAll x xs)"
   by (induct xs) (auto simp add: memb_def)
 
-lemma fdelete_raw_filter:
-  "delete_raw xs y = [x \<leftarrow> xs. x \<noteq> y]"
+lemma fremoveAll_filter:
+  "removeAll y xs = [x \<leftarrow> xs. x \<noteq> y]"
   by (induct xs) simp_all
 
 lemma fcard_raw_delete:
-  "fcard_raw (delete_raw xs y) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)"
-  by (simp add: fdelete_raw_filter fcard_raw_delete_one)
+  "fcard_raw (removeAll y xs) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)"
+  by (auto simp add: fcard_raw_def memb_def)
 
 lemma finter_raw_empty:
   "finter_raw l [] = []"
   by (induct l) (simp_all add: not_memb_nil)
 
-lemma set_cong:
-  shows "(x \<approx> y) = (set x = set y)"
-  by auto
-
 lemma inj_map_eq_iff:
   "inj f \<Longrightarrow> (map f l \<approx> map f m) = (l \<approx> m)"
   by (simp add: set_eq_iff[symmetric] inj_image_eq_iff)
@@ -878,7 +779,7 @@
   by (induct xs) (auto intro: list_eq2.intros)
 
 lemma cons_delete_list_eq2:
-  shows "list_eq2 (a # (delete_raw A a)) (if memb a A then A else a # A)"
+  shows "list_eq2 (a # (removeAll a A)) (if memb a A then A else a # A)"
   apply (induct A)
   apply (simp add: memb_def list_eq2_refl)
   apply (case_tac "memb a (aa # A)")
@@ -889,18 +790,18 @@
   apply (auto simp add: memb_def)[2]
   apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6))
   apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6))
-  apply (auto simp add: list_eq2_refl not_memb_delete_raw_ident)
+  apply (auto simp add: list_eq2_refl not_memb_removeAll_ident)
   done
 
 lemma memb_delete_list_eq2:
   assumes a: "memb e r"
-  shows "list_eq2 (e # delete_raw r e) r"
+  shows "list_eq2 (e # removeAll e r) r"
   using a cons_delete_list_eq2[of e r]
   by simp
 
-lemma delete_raw_rsp:
-  "xs \<approx> ys \<Longrightarrow> delete_raw xs x \<approx> delete_raw ys x"
-  by (simp add: memb_def[symmetric] memb_delete_raw)
+lemma removeAll_rsp:
+  "xs \<approx> ys \<Longrightarrow> removeAll x xs \<approx> removeAll x ys"
+  by (simp add: memb_def[symmetric])
 
 lemma list_eq2_equiv:
   "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)"
@@ -923,58 +824,27 @@
       case (Suc m)
       have b: "l \<approx> r" by fact
       have d: "fcard_raw l = Suc m" by fact
-      then have "\<exists>a. memb a l" by (rule fcard_raw_suc_memb)
+      then have "\<exists>a. memb a l" 
+	apply(simp add: fcard_raw_def memb_def)
+	apply(drule card_eq_SucD)
+	apply(blast)
+	done
       then obtain a where e: "memb a l" by auto
-      then have e': "memb a r" using list_eq.simps[simplified memb_def[symmetric], of l r] b by auto
-      have f: "fcard_raw (delete_raw l a) = m" using fcard_raw_delete[of l a] e d by simp
-      have g: "delete_raw l a \<approx> delete_raw r a" using delete_raw_rsp[OF b] by simp
-      have "list_eq2 (delete_raw l a) (delete_raw r a)" by (rule Suc.hyps[OF f g])
-      then have h: "list_eq2 (a # delete_raw l a) (a # delete_raw r a)" by (rule list_eq2.intros(5))
-      have i: "list_eq2 l (a # delete_raw l a)"
+      then have e': "memb a r" using list_eq.simps[simplified memb_def[symmetric], of l r] b 
+	unfolding memb_def by auto
+      have f: "fcard_raw (removeAll a l) = m" using fcard_raw_delete[of a l] e d by simp
+      have g: "removeAll a l \<approx> removeAll a r" using removeAll_rsp[OF b] by simp
+      have "list_eq2 (removeAll a l) (removeAll a r)" by (rule Suc.hyps[OF f g])
+      then have h: "list_eq2 (a # removeAll a l) (a # removeAll a r)" by (rule list_eq2.intros(5))
+      have i: "list_eq2 l (a # removeAll a l)"
         by (rule list_eq2.intros(3)[OF memb_delete_list_eq2[OF e]])
-      have "list_eq2 l (a # delete_raw r a)" by (rule list_eq2.intros(6)[OF i h])
+      have "list_eq2 l (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h])
       then show ?case using list_eq2.intros(6)[OF _ memb_delete_list_eq2[OF e']] by simp
     qed
     }
   then show "l \<approx> r \<Longrightarrow> list_eq2 l r" by blast
 qed
 
-text {* Set *}
-
-lemma sub_list_set: "sub_list xs ys = (set xs \<subseteq> set ys)"
-  by (metis rev_append set_append set_cong set_rev sub_list_append sub_list_append_left sub_list_def sub_list_not_eq subset_Un_eq)
-
-lemma sub_list_neq_set: "(sub_list xs ys \<and> \<not> list_eq xs ys) = (set xs \<subset> set ys)"
-  by (auto simp add: sub_list_set)
-
-lemma fcard_raw_set: "fcard_raw xs = card (set xs)"
-  by (induct xs) (auto simp add: insert_absorb memb_def card_insert_disjoint)
-
-lemma memb_set: "memb x xs = (x \<in> set xs)"
-  by (simp only: memb_def)
-
-lemma filter_set: "set (filter P xs) = P \<inter> (set xs)"
-  by (induct xs, simp)
-     (metis Int_insert_right_if0 Int_insert_right_if1 List.set.simps(2) filter.simps(2) mem_def)
-
-lemma delete_raw_set: "set (delete_raw xs x) = set xs - {x}"
-  by (induct xs) auto
-
-lemma inter_raw_set: "set (finter_raw xs ys) = set xs \<inter> set ys"
-  by (induct xs) (simp_all add: memb_def)
-
-lemma fminus_raw_set: "set (fminus_raw xs ys) = set xs - set ys"
-  by (induct ys arbitrary: xs)
-     (simp_all add: delete_raw_set, blast)
-
-text {* Raw theorems of ffilter *}
-
-lemma sub_list_filter: "sub_list (filter P xs) (filter Q xs) = (\<forall> x. memb x xs \<longrightarrow> P x \<longrightarrow> Q x)"
-unfolding sub_list_def memb_def by auto
-
-lemma list_eq_filter: "list_eq (filter P xs) (filter Q xs) = (\<forall>x. memb x xs \<longrightarrow> P x = Q x)"
-unfolding memb_def by auto
-
 text {* Lifted theorems *}
 
 lemma not_fin_fnil: "x |\<notin>| {||}"
@@ -1010,15 +880,16 @@
   shows "{|x|} = {|y|} \<longleftrightarrow> x = y"
   by (lifting singleton_list_eq)
 
-text {* fset_to_set *}
+
+text {* fset *}
 
-lemma fset_to_set_simps[simp]:
-  "fset_to_set {||} = ({} :: 'a set)"
-  "fset_to_set (finsert (h :: 'a) t) = insert h (fset_to_set t)"
+lemma fset_simps[simp]:
+  "fset {||} = ({} :: 'a set)"
+  "fset (finsert (h :: 'a) t) = insert h (fset t)"
   by (lifting set.simps)
 
-lemma in_fset_to_set:
-  "x \<in> fset_to_set S \<equiv> x |\<in>| S"
+lemma in_fset:
+  "x \<in> fset S \<equiv> x |\<in>| S"
   by (lifting memb_def[symmetric])
 
 lemma none_fin_fempty:
@@ -1026,48 +897,67 @@
   by (lifting none_memb_nil)
 
 lemma fset_cong:
-  "(S = T) = (fset_to_set S = fset_to_set T)"
-  by (lifting set_cong)
+  "(S = T) = (fset S = fset T)"
+  by (lifting list_eq.simps)
 
-text {* fcard *}
 
-lemma fcard_fempty [simp]:
-  shows "fcard {||} = 0"
-  by (lifting fcard_raw_nil)
+section {* fcard *}
 
 lemma fcard_finsert_if [simp]:
   shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))"
-  by (lifting fcard_raw_cons)
+  by (descending) (auto simp add: fcard_raw_def memb_def insert_absorb)
 
-lemma fcard_0: "(fcard S = 0) = (S = {||})"
-  by (lifting fcard_raw_0)
+lemma fcard_0[simp]: 
+  shows "fcard S = 0 \<longleftrightarrow> S = {||}"
+  by (descending) (simp add: fcard_raw_def)
+
+lemma fcard_fempty[simp]:
+  shows "fcard {||} = 0"
+  by (simp add: fcard_0)
 
 lemma fcard_1:
-  shows "(fcard S = 1) = (\<exists>x. S = {|x|})"
-  by (lifting fcard_raw_1)
+  shows "fcard S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})"
+  by (descending) (auto simp add: fcard_raw_def card_Suc_eq)
 
 lemma fcard_gt_0:
-  shows "x \<in> fset_to_set S \<Longrightarrow> 0 < fcard S"
-  by (lifting fcard_raw_gt_0)
-
+  shows "x \<in> fset S \<Longrightarrow> 0 < fcard S"
+  by (descending) (auto simp add: fcard_raw_def card_gt_0_iff)
+  
 lemma fcard_not_fin:
-  shows "(x |\<notin>| S) = (fcard (finsert x S) = Suc (fcard S))"
-  by (lifting fcard_raw_not_memb)
+  assumes a: "x |\<notin>| S" 
+  shows "fcard (finsert x S) = Suc (fcard S)"
+  using a
+  by (descending) (simp add: memb_def fcard_raw_def)
 
-lemma fcard_suc: "fcard S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = finsert x T \<and> fcard T = n"
-  by (lifting fcard_raw_suc)
+lemma fcard_suc: 
+  shows "fcard S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = finsert x T \<and> fcard T = n"
+  apply(descending)
+  apply(simp add: fcard_raw_def memb_def)
+  apply(drule card_eq_SucD)
+  apply(auto)
+  apply(rule_tac x="b" in exI)
+  apply(rule_tac x="removeAll b S" in exI)
+  apply(auto)
+  done
 
 lemma fcard_delete:
-  "fcard (fdelete S y) = (if y |\<in>| S then fcard S - 1 else fcard S)"
+  "fcard (fdelete y S) = (if y |\<in>| S then fcard S - 1 else fcard S)"
   by (lifting fcard_raw_delete)
 
-lemma fcard_suc_memb: "fcard A = Suc n \<Longrightarrow> \<exists>a. a |\<in>| A"
-  by (lifting fcard_raw_suc_memb)
+lemma fcard_suc_memb: 
+  shows "fcard A = Suc n \<Longrightarrow> \<exists>a. a |\<in>| A"
+  apply(descending)
+  apply(simp add: fcard_raw_def memb_def)
+  apply(drule card_eq_SucD)
+  apply(auto)
+  done
 
-lemma fin_fcard_not_0: "a |\<in>| A \<Longrightarrow> fcard A \<noteq> 0"
-  by (lifting memb_card_not_0)
+lemma fin_fcard_not_0: 
+  shows "a |\<in>| A \<Longrightarrow> fcard A \<noteq> 0"
+  by (descending) (auto simp add: fcard_raw_def memb_def)
 
-text {* funion *}
+
+section {* funion *}
 
 lemmas [simp] =
   sup_bot_left[where 'a="'a fset", standard]
@@ -1078,14 +968,15 @@
   by (lifting append.simps(2))
 
 lemma singleton_union_left:
-  "{|a|} |\<union>| S = finsert a S"
+  shows "{|a|} |\<union>| S = finsert a S"
   by simp
 
 lemma singleton_union_right:
-  "S |\<union>| {|a|} = finsert a S"
+  shows "S |\<union>| {|a|} = finsert a S"
   by (subst sup.commute) simp
 
-section {* Induction and Cases rules for finite sets *}
+
+section {* Induction and Cases rules for fsets *}
 
 lemma fset_strong_cases:
   obtains "xs = {||}"
@@ -1141,7 +1032,8 @@
   then show "P (finsert x zs)" using b h by simp
 qed
 
-text {* fmap *}
+
+section {* fmap *}
 
 lemma fmap_simps[simp]:
   "fmap (f :: 'a \<Rightarrow> 'b) {||} = {||}"
@@ -1149,7 +1041,7 @@
   by (lifting map.simps)
 
 lemma fmap_set_image:
-  "fset_to_set (fmap f S) = f ` (fset_to_set S)"
+  "fset (fmap f S) = f ` (fset S)"
   by (induct S) simp_all
 
 lemma inj_fmap_eq_iff:
@@ -1163,76 +1055,88 @@
   "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T"
   by (lifting memb_append)
 
-text {* to_set *}
+
+section {* fset *}
 
-lemma fin_set: "(x |\<in>| xs) = (x \<in> fset_to_set xs)"
-  by (lifting memb_set)
+lemma fin_set: 
+  shows "x |\<in>| xs \<longleftrightarrow> x \<in> fset xs"
+  by (lifting memb_def)
 
-lemma fnotin_set: "(x |\<notin>| xs) = (x \<notin> fset_to_set xs)"
+lemma fnotin_set: 
+  shows "x |\<notin>| xs \<longleftrightarrow> x \<notin> fset xs"
   by (simp add: fin_set)
 
-lemma fcard_set: "fcard xs = card (fset_to_set xs)"
-  by (lifting fcard_raw_set)
+lemma fcard_set: 
+  shows "fcard xs = card (fset xs)"
+  by (lifting fcard_raw_def)
 
-lemma fsubseteq_set: "(xs |\<subseteq>| ys) = (fset_to_set xs \<subseteq> fset_to_set ys)"
-  by (lifting sub_list_set)
+lemma fsubseteq_set: 
+  shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys"
+  by (lifting sub_list_def)
 
-lemma fsubset_set: "(xs |\<subset>| ys) = (fset_to_set xs \<subset> fset_to_set ys)"
-  unfolding less_fset by (lifting sub_list_neq_set)
+lemma fsubset_set: 
+  shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys"
+  unfolding less_fset_def 
+  by (descending) (auto simp add: sub_list_def)
 
-lemma ffilter_set: "fset_to_set (ffilter P xs) = P \<inter> fset_to_set xs"
-  by (lifting filter_set)
+lemma ffilter_set [simp]: 
+  shows "fset (ffilter P xs) = P \<inter> fset xs"
+  by (descending) (auto simp add: mem_def)
 
-lemma fdelete_set: "fset_to_set (fdelete xs x) = fset_to_set xs - {x}"
-  by (lifting delete_raw_set)
+lemma fdelete_set [simp]: 
+  shows "fset (fdelete x xs) = fset xs - {x}"
+  by (lifting set_removeAll)
 
-lemma inter_set: "fset_to_set (xs |\<inter>| ys) = fset_to_set xs \<inter> fset_to_set ys"
-  by (lifting inter_raw_set)
+lemma finter_set [simp]: 
+  shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys"
+  by (lifting set_finter_raw)
 
-lemma union_set: "fset_to_set (xs |\<union>| ys) = fset_to_set xs \<union> fset_to_set ys"
+lemma funion_set [simp]: 
+  shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys"
   by (lifting set_append)
 
-lemma fminus_set: "fset_to_set (xs - ys) = fset_to_set xs - fset_to_set ys"
-  by (lifting fminus_raw_set)
-
-lemmas fset_to_set_trans =
-  fin_set fnotin_set fcard_set fsubseteq_set fsubset_set
-  inter_set union_set ffilter_set fset_to_set_simps
-  fset_cong fdelete_set fmap_set_image fminus_set
+lemma fminus_set [simp]: 
+  shows "fset (xs - ys) = fset xs - fset ys"
+  by (lifting set_fminus_raw)
 
 
-text {* ffold *}
+
+section {* ffold *}
 
-lemma ffold_nil: "ffold f z {||} = z"
+lemma ffold_nil: 
+  shows "ffold f z {||} = z"
   by (lifting ffold_raw.simps(1)[where 'a="'b" and 'b="'a"])
 
 lemma ffold_finsert: "ffold f z (finsert a A) =
   (if rsp_fold f then if a |\<in>| A then ffold f z A else f a (ffold f z A) else z)"
-  by (lifting ffold_raw.simps(2)[where 'a="'b" and 'b="'a"])
+  by (descending) (simp add: memb_def)
 
 lemma fin_commute_ffold:
-  "\<lbrakk>rsp_fold f; h |\<in>| b\<rbrakk> \<Longrightarrow> ffold f z b = f h (ffold f z (fdelete b h))"
-  by (lifting memb_commute_ffold_raw)
+  "\<lbrakk>rsp_fold f; h |\<in>| b\<rbrakk> \<Longrightarrow> ffold f z b = f h (ffold f z (fdelete h b))"
+  by (descending) (simp add: memb_def memb_commute_ffold_raw)
 
-text {* fdelete *}
+
+
+section {* fdelete *}
 
 lemma fin_fdelete:
-  shows "x |\<in>| fdelete S y \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y"
-  by (lifting memb_delete_raw)
+  shows "x |\<in>| fdelete y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y"
+  by (descending) (simp add: memb_def)
 
 lemma fin_fdelete_ident:
-  shows "x |\<notin>| fdelete S x"
-  by (lifting memb_delete_raw_ident)
+  shows "x |\<notin>| fdelete x S"
+  by (lifting memb_removeAll_ident)
 
 lemma not_memb_fdelete_ident:
-  shows "x |\<notin>| S \<Longrightarrow> fdelete S x = S"
-  by (lifting not_memb_delete_raw_ident)
+  shows "x |\<notin>| S \<Longrightarrow> fdelete x S = S"
+  by (lifting not_memb_removeAll_ident)
 
 lemma fset_fdelete_cases:
-  shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = finsert x (fdelete S x))"
-  by (lifting fset_raw_delete_raw_cases)
+  shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = finsert x (fdelete x S))"
+  by (lifting fset_raw_removeAll_cases)
 
-text {* inter *}
+
+section {* finter *}
 
 lemma finter_empty_l: "({||} |\<inter>| S) = {||}"
   by (lifting finter_raw.simps(1))
@@ -1241,39 +1145,44 @@
   by (lifting finter_raw_empty)
 
 lemma finter_finsert:
-  "finsert x S |\<inter>| T = (if x |\<in>| T then finsert x (S |\<inter>| T) else S |\<inter>| T)"
-  by (lifting finter_raw.simps(2))
+  shows "finsert x S |\<inter>| T = (if x |\<in>| T then finsert x (S |\<inter>| T) else S |\<inter>| T)"
+  by (descending) (simp add: memb_def)
 
 lemma fin_finter:
-  "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T"
-  by (lifting memb_finter_raw)
+  shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T"
+  by (descending) (simp add: memb_def)
 
 lemma fsubset_finsert:
-  "(finsert x xs |\<subseteq>| ys) = (x |\<in>| ys \<and> xs |\<subseteq>| ys)"
+  shows "(finsert x xs |\<subseteq>| ys) = (x |\<in>| ys \<and> xs |\<subseteq>| ys)"
   by (lifting sub_list_cons)
 
-lemma "xs |\<subseteq>| ys \<equiv> \<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys"
-  by (lifting sub_list_def[simplified memb_def[symmetric]])
+lemma 
+  shows "xs |\<subseteq>| ys \<equiv> \<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys"
+  by (descending) (auto simp add: sub_list_def memb_def)
 
-lemma fsubset_fin: "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)"
-by (rule meta_eq_to_obj_eq)
-   (lifting sub_list_def[simplified memb_def[symmetric]])
+lemma fsubset_fin: 
+  shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)"
+  by (descending) (auto simp add: sub_list_def memb_def)
 
-lemma fminus_fin: "(x |\<in>| xs - ys) = (x |\<in>| xs \<and> x |\<notin>| ys)"
-  by (lifting fminus_raw_memb)
+lemma fminus_fin: 
+  shows "(x |\<in>| xs - ys) = (x |\<in>| xs \<and> x |\<notin>| ys)"
+  by (descending) (simp add: memb_def)
 
-lemma fminus_red: "finsert x xs - ys = (if x |\<in>| ys then xs - ys else finsert x (xs - ys))"
-  by (lifting fminus_raw_red)
+lemma fminus_red: 
+  shows "finsert x xs - ys = (if x |\<in>| ys then xs - ys else finsert x (xs - ys))"
+  by (descending) (auto simp add: memb_def)
 
-lemma fminus_red_fin[simp]: "x |\<in>| ys \<Longrightarrow> finsert x xs - ys = xs - ys"
+lemma fminus_red_fin[simp]: 
+  shows "x |\<in>| ys \<Longrightarrow> finsert x xs - ys = xs - ys"
   by (simp add: fminus_red)
 
-lemma fminus_red_fnotin[simp]: "x |\<notin>| ys \<Longrightarrow> finsert x xs - ys = finsert x (xs - ys)"
+lemma fminus_red_fnotin[simp]: 
+  shows "x |\<notin>| ys \<Longrightarrow> finsert x xs - ys = finsert x (xs - ys)"
   by (simp add: fminus_red)
 
 lemma fset_eq_iff:
   "(S = T) = (\<forall>x. (x |\<in>| S) = (x |\<in>| T))"
-  by (lifting list_eq.simps[simplified memb_def[symmetric]])
+  by (descending) (auto simp add: memb_def)
 
 (* We cannot write it as "assumes .. shows" since Isabelle changes
    the quantifiers to schematic variables and reintroduces them in
@@ -1300,7 +1209,8 @@
   using assms
   by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]])
 
-text {* concat *}
+
+section {* fconcat *}
 
 lemma fconcat_empty:
   shows "fconcat {||} = {||}"
@@ -1310,110 +1220,131 @@
   shows "fconcat (finsert x S) = x |\<union>| fconcat S"
   by (lifting concat.simps(2))
 
-text {* ffilter *}
+
+section {* ffilter *}
 
-lemma subseteq_filter: "ffilter P xs <= ffilter Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)"
-by (lifting sub_list_filter)
+lemma subseteq_filter: 
+  shows "ffilter P xs <= ffilter Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)"
+  by  (descending) (auto simp add: memb_def sub_list_def)
 
-lemma eq_ffilter: "(ffilter P xs = ffilter Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)"
-by (lifting list_eq_filter)
-
+lemma eq_ffilter: 
+  shows "(ffilter P xs = ffilter Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)"
+  by (descending) (auto simp add: memb_def)
+  
 lemma subset_ffilter: 
   "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> ffilter P xs < ffilter Q xs"
-unfolding less_fset by (auto simp add: subseteq_filter eq_ffilter)
+unfolding less_fset_def by (auto simp add: subseteq_filter eq_ffilter)
+
 
 section {* lemmas transferred from Finite_Set theory *}
 
 text {* finiteness for finite sets holds *}
-lemma finite_fset: "finite (fset_to_set S)"
+lemma finite_fset [simp]: 
+  shows "finite (fset S)"
   by (induct S) auto
 
-lemma fset_choice: "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)"
-  unfolding fset_to_set_trans
-  by (rule finite_set_choice[simplified Ball_def, OF finite_fset])
+lemma fset_choice: 
+  shows "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)"
+  apply(descending)
+  apply(simp add: memb_def)
+  apply(rule finite_set_choice[simplified Ball_def])
+  apply(simp_all)
+  done
 
-lemma fsubseteq_fnil: "xs |\<subseteq>| {||} = (xs = {||})"
-  unfolding fset_to_set_trans
-  by (rule subset_empty)
+lemma fsubseteq_fempty: 
+  shows "xs |\<subseteq>| {||} = (xs = {||})"
+  by (metis finter_empty_r le_iff_inf)
 
-lemma not_fsubset_fnil: "\<not> xs |\<subset>| {||}"
-  unfolding fset_to_set_trans
-  by (rule not_psubset_empty)
-
-lemma fcard_mono: "xs |\<subseteq>| ys \<Longrightarrow> fcard xs \<le> fcard ys"
-  unfolding fset_to_set_trans
+lemma not_fsubset_fnil: 
+  shows "\<not> xs |\<subset>| {||}"
+  by (metis fset_simps(1) fsubset_set not_psubset_empty)
+  
+lemma fcard_mono: 
+  shows "xs |\<subseteq>| ys \<Longrightarrow> fcard xs \<le> fcard ys"
+  unfolding fcard_set fsubseteq_set
   by (rule card_mono[OF finite_fset])
 
-lemma fcard_fseteq: "xs |\<subseteq>| ys \<Longrightarrow> fcard ys \<le> fcard xs \<Longrightarrow> xs = ys"
-  unfolding fset_to_set_trans
-  by (rule card_seteq[OF finite_fset])
+lemma fcard_fseteq: 
+  shows "xs |\<subseteq>| ys \<Longrightarrow> fcard ys \<le> fcard xs \<Longrightarrow> xs = ys"
+  unfolding fcard_set fsubseteq_set
+  by (simp add: card_seteq[OF finite_fset] fset_cong)
 
-lemma psubset_fcard_mono: "xs |\<subset>| ys \<Longrightarrow> fcard xs < fcard ys"
-  unfolding fset_to_set_trans
+lemma psubset_fcard_mono: 
+  shows "xs |\<subset>| ys \<Longrightarrow> fcard xs < fcard ys"
+  unfolding fcard_set fsubset_set
   by (rule psubset_card_mono[OF finite_fset])
 
-lemma fcard_funion_finter: "fcard xs + fcard ys = fcard (xs |\<union>| ys) + fcard (xs |\<inter>| ys)"
-  unfolding fset_to_set_trans
+lemma fcard_funion_finter: 
+  shows "fcard xs + fcard ys = fcard (xs |\<union>| ys) + fcard (xs |\<inter>| ys)"
+  unfolding fcard_set funion_set finter_set
   by (rule card_Un_Int[OF finite_fset finite_fset])
 
-lemma fcard_funion_disjoint: "xs |\<inter>| ys = {||} \<Longrightarrow> fcard (xs |\<union>| ys) = fcard xs + fcard ys"
-  unfolding fset_to_set_trans
-  by (rule card_Un_disjoint[OF finite_fset finite_fset])
+lemma fcard_funion_disjoint: 
+  shows "xs |\<inter>| ys = {||} \<Longrightarrow> fcard (xs |\<union>| ys) = fcard xs + fcard ys"
+  unfolding fcard_set funion_set 
+  apply (rule card_Un_disjoint[OF finite_fset finite_fset])
+  by (metis finter_set fset_simps(1))
 
-lemma fcard_delete1_less: "x |\<in>| xs \<Longrightarrow> fcard (fdelete xs x) < fcard xs"
-  unfolding fset_to_set_trans
+lemma fcard_delete1_less: 
+  shows "x |\<in>| xs \<Longrightarrow> fcard (fdelete x xs) < fcard xs"
+  unfolding fcard_set fin_set fdelete_set 
   by (rule card_Diff1_less[OF finite_fset])
 
-lemma fcard_delete2_less: "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> fcard (fdelete (fdelete xs x) y) < fcard xs"
-  unfolding fset_to_set_trans
+lemma fcard_delete2_less: 
+  shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> fcard (fdelete y (fdelete x xs)) < fcard xs"
+  unfolding fcard_set fdelete_set fin_set
   by (rule card_Diff2_less[OF finite_fset])
 
-lemma fcard_delete1_le: "fcard (fdelete xs x) <= fcard xs"
-  unfolding fset_to_set_trans
+lemma fcard_delete1_le: 
+  shows "fcard (fdelete x xs) <= fcard xs"
+  unfolding fdelete_set fcard_set
   by (rule card_Diff1_le[OF finite_fset])
 
-lemma fcard_psubset: "ys |\<subseteq>| xs \<Longrightarrow> fcard ys < fcard xs \<Longrightarrow> ys |\<subset>| xs"
-  unfolding fset_to_set_trans
+lemma fcard_psubset: 
+  shows "ys |\<subseteq>| xs \<Longrightarrow> fcard ys < fcard xs \<Longrightarrow> ys |\<subset>| xs"
+  unfolding fcard_set fsubseteq_set fsubset_set
   by (rule card_psubset[OF finite_fset])
 
-lemma fcard_fmap_le: "fcard (fmap f xs) \<le> fcard xs"
-  unfolding fset_to_set_trans
+lemma fcard_fmap_le: 
+  shows "fcard (fmap f xs) \<le> fcard xs"
+  unfolding fcard_set  fmap_set_image
   by (rule card_image_le[OF finite_fset])
 
-lemma fin_fminus_fnotin: "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S"
-  unfolding fset_to_set_trans
+lemma fin_fminus_fnotin: 
+  shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S"
+  unfolding fin_set fminus_set
   by blast
 
-lemma fin_fnotin_fminus: "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S"
-  unfolding fset_to_set_trans
+lemma fin_fnotin_fminus: 
+  shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S"
+  unfolding fin_set fminus_set
   by blast
 
-lemma fin_mdef: "x |\<in>| F = ((x |\<notin>| (F - {|x|})) & (F = finsert x (F - {|x|})))"
-  unfolding fset_to_set_trans
+lemma fin_mdef: 
+  shows "x |\<in>| F = ((x |\<notin>| (F - {|x|})) & (F = finsert x (F - {|x|})))"
+  unfolding fin_set fset_simps fset_cong fminus_set
   by blast
 
 lemma fcard_fminus_finsert[simp]:
   assumes "a |\<in>| A" and "a |\<notin>| B"
-  shows "fcard(A - finsert a B) = fcard(A - B) - 1"
-  using assms unfolding fset_to_set_trans
-  by (rule card_Diff_insert[OF finite_fset])
+  shows "fcard (A - finsert a B) = fcard (A - B) - 1"
+  using assms 
+  unfolding fin_set fcard_set fminus_set
+  by (simp add: card_Diff_insert[OF finite_fset])
 
 lemma fcard_fminus_fsubset:
   assumes "B |\<subseteq>| A"
   shows "fcard (A - B) = fcard A - fcard B"
-  using assms unfolding fset_to_set_trans
+  using assms 
+  unfolding fsubseteq_set fcard_set fminus_set
   by (rule card_Diff_subset[OF finite_fset])
 
 lemma fcard_fminus_subset_finter:
-  "fcard (A - B) = fcard A - fcard (A |\<inter>| B)"
-  unfolding fset_to_set_trans
-  by (rule card_Diff_subset_Int) (fold inter_set, rule finite_fset)
+  shows "fcard (A - B) = fcard A - fcard (A |\<inter>| B)"
+  using assms 
+  unfolding finter_set fcard_set fminus_set
+  by (rule card_Diff_subset_Int) (simp)
 
-lemma ball_reg_right_unfolded: "(\<forall>x. R x \<longrightarrow> P x \<longrightarrow> Q x) \<longrightarrow> (All P \<longrightarrow> Ball R Q)"
-apply rule
-apply (rule ball_reg_right)
-apply auto
-done
 
 lemma list_all2_refl:
   assumes q: "equivp R"
--- a/Nominal/Nominal2_FSet.thy	Wed Oct 13 22:55:58 2010 +0100
+++ b/Nominal/Nominal2_FSet.thy	Thu Oct 14 04:14:22 2010 +0100
@@ -24,10 +24,8 @@
 instance 
 proof
   fix x :: "'a fset" and p q :: "perm"
-  show "0 \<bullet> x = x"
-    by (lifting permute_zero [where 'a="'a list"])
-  show "(p + q) \<bullet> x = p \<bullet> q \<bullet> x"
-    by (lifting permute_plus [where 'a="'a list"])
+  show "0 \<bullet> x = x" by (descending) (simp)
+  show "(p + q) \<bullet> x = p \<bullet> q \<bullet> x" by (descending) (simp)
 qed
 
 end
@@ -42,16 +40,12 @@
   shows "p \<bullet> (fmap f S) = fmap (p \<bullet> f) (p \<bullet> S)"
   by (lifting map_eqvt)
 
-lemma fset_to_set_eqvt [eqvt]: 
-  shows "p \<bullet> (fset_to_set S) = fset_to_set (p \<bullet> S)"
+lemma fset_eqvt[eqvt]: 
+  shows "p \<bullet> (fset S) = fset (p \<bullet> S)"
   by (lifting set_eqvt)
 
-lemma fin_fset_to_set[simp]:
-  shows "finite (fset_to_set S)"
-  by (induct S) (simp_all)
-
-lemma supp_fset_to_set:
-  shows "supp (fset_to_set S) = supp S"
+lemma supp_fset:
+  shows "supp (fset S) = supp S"
   unfolding supp_def
   by (perm_simp) (simp add: fset_cong)
 
@@ -62,11 +56,12 @@
 
 lemma supp_finsert:
   fixes x::"'a::fs"
+  and   S::"'a fset"
   shows "supp (finsert x S) = supp x \<union> supp S"
-  apply(subst supp_fset_to_set[symmetric])
-  apply(simp add: supp_fset_to_set)
+  apply(subst supp_fset[symmetric])
+  apply(simp add: supp_fset)
   apply(simp add: supp_of_fin_insert)
-  apply(simp add: supp_fset_to_set)
+  apply(simp add: supp_fset)
   done
 
 
@@ -93,7 +88,7 @@
 
 lemma supp_at_fset:
   fixes S::"('a::at_base) fset"
-  shows "supp S = fset_to_set (fmap atom S)"
+  shows "supp S = fset (fmap atom S)"
   apply (induct S)
   apply (simp add: supp_fempty)
   apply (simp add: supp_finsert)
@@ -102,7 +97,7 @@
 
 lemma fresh_star_atom:
   fixes a::"'a::at_base"
-  shows "fset_to_set S \<sharp>* a \<Longrightarrow> atom a \<sharp> fset_to_set S"
+  shows "fset S \<sharp>* a \<Longrightarrow> atom a \<sharp> fset S"
   apply (induct S)
   apply (simp add: fresh_set_empty)
   apply simp
--- a/Nominal/nominal_dt_rawfuns.ML	Wed Oct 13 22:55:58 2010 +0100
+++ b/Nominal/nominal_dt_rawfuns.ML	Thu Oct 14 04:14:22 2010 +0100
@@ -99,9 +99,9 @@
     val ty = fastype_of t;
     val atom_ty = dest_fsetT ty --> @{typ "atom"};
     val fmap_ty = atom_ty --> ty --> @{typ "atom fset"};
-    val fset_to_set = @{term "fset_to_set :: atom fset => atom set"}
+    val fset = @{term "fset :: atom fset => atom set"}
   in
-    fset_to_set $ (Const (@{const_name fmap}, fmap_ty) $ Const (@{const_name atom}, atom_ty) $ t)
+    fset $ (Const (@{const_name fmap}, fmap_ty) $ Const (@{const_name atom}, atom_ty) $ t)
   end
 
   fun mk_atom_list t =