diff -r f374ffd50c7c -r ac8cb569a17b Nominal/FSet.thy --- a/Nominal/FSet.thy Wed Apr 14 10:29:34 2010 +0200 +++ b/Nominal/FSet.thy Wed Apr 14 10:29:56 2010 +0200 @@ -245,9 +245,9 @@ "x # x # xs \ x # xs" by auto -lemma none_mem_nil: - "(\x. x \ set xs) = (xs \ [])" - by simp +lemma none_memb_nil: + "(\x. \ memb x xs) = (xs \ [])" + by (simp add: memb_def) lemma fset_raw_strong_cases: "(xs = []) \ (\x ys. ((\ memb x ys) \ (xs \ x # ys)))" @@ -278,6 +278,10 @@ "memb x (delete_raw xs y) = (memb x xs \ x \ y)" by (induct xs arbitrary: x y) (auto simp add: memb_def) +lemma delete_raw_rsp: + "l \ r \ delete_raw l x \ delete_raw r x" + by (simp add: memb_def[symmetric] memb_delete_raw) + lemma [quot_respect]: "(op \ ===> op = ===> op \) delete_raw delete_raw" by (simp add: memb_def[symmetric] memb_delete_raw) @@ -339,7 +343,7 @@ lemma ffold_raw_rsp_pre: "\e. memb e a = memb e b \ ffold_raw f z a = ffold_raw f z b" apply (induct a arbitrary: b) - apply (simp add: hd_in_set memb_absorb memb_def none_mem_nil) + apply (simp add: hd_in_set memb_absorb memb_def none_memb_nil) apply (simp add: ffold_raw.simps) apply (rule conjI) apply (rule_tac [!] impI) @@ -479,15 +483,13 @@ "fset_to_set (finsert (h :: 'a) t) = insert h (fset_to_set t)" by (lifting set.simps) -thm memb_def[symmetric, THEN meta_eq_to_obj_eq] - lemma in_fset_to_set: "x \ fset_to_set S \ x |\| S" by (lifting memb_def[symmetric]) lemma none_fin_fempty: - "(\x. x \ fset_to_set S) = (S = {||})" - by (lifting none_mem_nil) + "(\x. x |\| S) = (S = {||})" + by (lifting none_memb_nil) lemma fset_cong: "(fset_to_set S = fset_to_set T) = (S = T)"