diff -r 509a0ccc4f32 -r 4465723e35e7 Nominal/FSet.thy --- a/Nominal/FSet.thy Wed Apr 14 07:34:03 2010 +0200 +++ b/Nominal/FSet.thy Wed Apr 14 07:57:55 2010 +0200 @@ -131,13 +131,11 @@ lemma fcard_raw_not_memb: fixes x :: "'a" - fixes xs :: "'a list" shows "\(memb x xs) \ (fcard_raw (x # xs) = Suc (fcard_raw xs))" by auto lemma fcard_raw_suc: fixes xs :: "'a list" - fixes n :: "nat" assumes c: "fcard_raw xs = Suc n" shows "\x ys. \(memb x ys) \ xs \ (x # ys) \ fcard_raw ys = n" unfolding memb_def @@ -230,18 +228,9 @@ text {* raw section *} -lemma map_rsp_aux: - assumes a: "a \ b" - shows "map f a \ map f b" - using a - apply(induct a arbitrary: b) - apply(auto) - apply(metis rev_image_eqI) - done - lemma map_rsp[quot_respect]: shows "(op = ===> op \ ===> op \) map map" - by (auto simp add: map_rsp_aux) + by auto lemma cons_left_comm: "x # y # xs \ y # x # xs" @@ -256,14 +245,14 @@ by simp lemma fset_raw_strong_cases: - "(xs = []) \ (\x ys. ((x \ set ys) \ (xs \ x # ys)))" + "(xs = []) \ (\x ys. ((\ memb x ys) \ (xs \ x # ys)))" apply (induct xs) apply (simp) apply (rule disjI2) apply (erule disjE) apply (rule_tac x="a" in exI) apply (rule_tac x="[]" in exI) - apply (simp) + apply (simp add: memb_def) apply (erule exE)+ apply (case_tac "x = a") apply (rule_tac x="a" in exI) @@ -271,7 +260,7 @@ apply (simp) apply (rule_tac x="x" in exI) apply (rule_tac x="a # ys" in exI) - apply (auto) + apply (auto simp add: memb_def) done fun @@ -386,8 +375,8 @@ by (induct l) (simp_all add: not_memb_nil) lemma memb_finter_raw: - "memb e (finter_raw l r) = (memb e l \ memb e r)" - apply (induct l) + "memb x (finter_raw xs ys) = (memb x xs \ memb x ys)" + apply (induct xs) apply (simp add: not_memb_nil) apply (simp add: finter_raw.simps) apply (simp add: memb_cons_iff) @@ -419,12 +408,12 @@ is "finter_raw" lemma funion_sym_pre: - "a @ b \ b @ a" + "xs @ ys \ ys @ xs" by auto lemma append_rsp[quot_respect]: shows "(op \ ===> op \ ===> op \) op @ op @" - by (auto) + by auto lemma set_cong: shows "(set x = set y) = (x \ y)" @@ -467,11 +456,11 @@ by (lifting nil_not_cons) lemma finsert_left_comm: - "finsert a (finsert b S) = finsert b (finsert a S)" + "finsert x (finsert y S) = finsert y (finsert x S)" by (lifting cons_left_comm) lemma finsert_left_idem: - "finsert a (finsert a S) = finsert a S" + "finsert x (finsert x S) = finsert x S" by (lifting cons_left_idem) lemma fsingleton_eq[simp]: @@ -485,16 +474,18 @@ "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 xs \ x |\| xs" + "x \ fset_to_set S \ x |\| S" by (lifting memb_def[symmetric]) lemma none_fin_fempty: - "(\a. a \ fset_to_set A) = (A = {||})" + "(\x. x \ fset_to_set S) = (S = {||})" by (lifting none_mem_nil) lemma fset_cong: - "(fset_to_set x = fset_to_set y) = (x = y)" + "(fset_to_set S = fset_to_set T) = (S = T)" by (lifting set_cong) text {* fcard *} @@ -507,46 +498,46 @@ shows "fcard (finsert x S) = (if x |\| S then fcard S else Suc (fcard S))" by (lifting fcard_raw_cons) -lemma fcard_0: "(fcard a = 0) = (a = {||})" +lemma fcard_0: "(fcard S = 0) = (S = {||})" by (lifting fcard_raw_0) lemma fcard_1: - fixes xs::"'b fset" - shows "(fcard xs = 1) = (\x. xs = {|x|})" + fixes S::"'b fset" + shows "(fcard S = 1) = (\x. S = {|x|})" by (lifting fcard_raw_1) -lemma fcard_gt_0: "x \ fset_to_set xs \ 0 < fcard xs" +lemma fcard_gt_0: "x \ fset_to_set S \ 0 < fcard S" by (lifting fcard_raw_gt_0) -lemma fcard_not_fin: "(x |\| xs) = (fcard (finsert x xs) = Suc (fcard xs))" +lemma fcard_not_fin: "(x |\| S) = (fcard (finsert x S) = Suc (fcard S))" by (lifting fcard_raw_not_memb) -lemma fcard_suc: "fcard xs = Suc n \ \a ys. a |\| ys \ xs = finsert a ys \ fcard ys = n" +lemma fcard_suc: "fcard S = Suc n \ \x T. x |\| T \ S = finsert x T \ fcard T = n" by (lifting fcard_raw_suc) lemma fcard_delete: - "fcard (fdelete xs y) = (if y |\| xs then fcard xs - 1 else fcard xs)" + "fcard (fdelete S y) = (if y |\| S then fcard S - 1 else fcard S)" by (lifting fcard_raw_delete) text {* funion *} lemma funion_simps[simp]: - "{||} |\| ys = ys" - "finsert x xs |\| ys = finsert x (xs |\| ys)" + "{||} |\| S = S" + "finsert x S |\| T = finsert x (S |\| T)" by (lifting append.simps) lemma funion_sym: - "a |\| b = b |\| a" + "S |\| T = T |\| S" by (lifting funion_sym_pre) lemma funion_assoc: - "x |\| xa |\| xb = x |\| (xa |\| xb)" + "S |\| T |\| U = S |\| (T |\| U)" by (lifting append_assoc) section {* Induction and Cases rules for finite sets *} lemma fset_strong_cases: - "X = {||} \ (\a Y. a \ fset_to_set Y \ X = finsert a Y)" + "S = {||} \ (\x T. x |\| T \ S = finsert x T)" by (lifting fset_raw_strong_cases) lemma fset_exhaust[case_names fempty finsert, cases type: fset]: @@ -596,24 +587,22 @@ lemma fmap_simps[simp]: "fmap (f :: 'a \ 'b) {||} = {||}" - "fmap f (finsert x xs) = finsert (f x) (fmap f xs)" + "fmap f (finsert x S) = finsert (f x) (fmap f S)" by (lifting map.simps) lemma fmap_set_image: - "fset_to_set (fmap f fs) = f ` (fset_to_set fs)" - apply (induct fs) - apply (simp_all) -done + "fset_to_set (fmap f S) = f ` (fset_to_set S)" + by (induct S) (simp_all) lemma inj_fmap_eq_iff: - "inj f \ (fmap f l = fmap f m) = (l = m)" + "inj f \ (fmap f S = fmap f T) = (S = T)" by (lifting inj_map_eq_iff) -lemma fmap_funion: "fmap f (a |\| b) = fmap f a |\| fmap f b" +lemma fmap_funion: "fmap f (S |\| T) = fmap f S |\| fmap f T" by (lifting map_append) lemma fin_funion: - "(e |\| l |\| r) = (e |\| l \ e |\| r)" + "x |\| S |\| T \ x |\| S \ x |\| T" by (lifting memb_append) text {* ffold *} @@ -631,37 +620,40 @@ text {* fdelete *} -lemma fin_fdelete: "(x |\| fdelete A a) = (x |\| A \ x \ a)" +lemma fin_fdelete: + shows "x |\| fdelete S y \ x |\| S \ x \ y" by (lifting memb_delete_raw) -lemma fin_fdelete_ident: "a |\| fdelete A a" +lemma fin_fdelete_ident: + shows "x |\| fdelete S x" by (lifting memb_delete_raw_ident) -lemma not_memb_fdelete_ident: "b |\| A \ fdelete A b = A" +lemma not_memb_fdelete_ident: + shows "x |\| S \ fdelete S x = S" by (lifting not_memb_delete_raw_ident) lemma fset_fdelete_cases: - "X = {||} \ (\a. a |\| X \ X = finsert a (fdelete X a))" + shows "S = {||} \ (\x. x |\| S \ S = finsert x (fdelete S x))" by (lifting fset_raw_delete_raw_cases) text {* inter *} -lemma finter_empty_l: "({||} |\| r) = {||}" +lemma finter_empty_l: "({||} |\| S) = {||}" by (lifting finter_raw.simps(1)) -lemma finter_empty_r: "(l |\| {||}) = {||}" +lemma finter_empty_r: "(S |\| {||}) = {||}" by (lifting finter_raw_empty) lemma finter_finsert: - "(finsert h t |\| l) = (if h |\| l then finsert h (t |\| l) else t |\| l)" + "finsert x S |\| T = (if x |\| T then finsert x (S |\| T) else S |\| T)" by (lifting finter_raw.simps(2)) lemma fin_finter: - "(e |\| (l |\| r)) = (e |\| l \ e |\| r)" + "x |\| (S |\| T) \ x |\| S \ x |\| T" by (lifting memb_finter_raw) lemma expand_fset_eq: - "(xs = ys) = (\x. (x |\| xs) = (x |\| ys))" + "(S = T) = (\x. (x |\| S) = (x |\| T))" by (lifting list_eq.simps[simplified memb_def[symmetric]])