# HG changeset patch # User Cezary Kaliszyk # Date 1271683999 -7200 # Node ID 4df853f5879f89adaf7f69f4e50f7257c9048b40 # Parent 6c5b5ec53a0b43d1d5485a20c6374801a4b324db reorder diff -r 6c5b5ec53a0b -r 4df853f5879f Nominal/FSet.thy --- a/Nominal/FSet.thy Mon Apr 19 15:28:57 2010 +0200 +++ b/Nominal/FSet.thy Mon Apr 19 15:33:19 2010 +0200 @@ -29,25 +29,6 @@ where "sub_list xs ys \ (\x. x \ set xs \ x \ set ys)" -lemma sub_list_cons: - "sub_list (x # xs) ys = (memb x ys \ sub_list xs ys)" - by (auto simp add: memb_def sub_list_def) - -lemma nil_not_cons: - shows "\ ([] \ x # xs)" - and "\ (x # xs \ [])" - by auto - -lemma sub_list_rsp1: "\xs \ ys\ \ sub_list xs zs = sub_list ys zs" - by (simp add: sub_list_def) - -lemma sub_list_rsp2: "\xs \ ys\ \ sub_list zs xs = sub_list zs ys" - by (simp add: sub_list_def) - -lemma [quot_respect]: - "(op \ ===> op \ ===> op =) sub_list sub_list" - by (auto simp only: fun_rel_def sub_list_rsp1 sub_list_rsp2) - quotient_type 'a fset = "'a list" / "list_eq" by (rule list_eq_equivp) @@ -95,6 +76,11 @@ section {* Augmenting an fset -- @{const finsert} *} +lemma nil_not_cons: + shows "\ ([] \ x # xs)" + and "\ (x # xs \ [])" + by auto + lemma not_memb_nil: shows "\ memb x []" by (simp add: memb_def) @@ -138,6 +124,22 @@ is "op @" +section {* sub_list *} + +lemma sub_list_cons: + "sub_list (x # xs) ys = (memb x ys \ sub_list xs ys)" + by (auto simp add: memb_def sub_list_def) + +lemma sub_list_rsp1: "\xs \ ys\ \ sub_list xs zs = sub_list ys zs" + by (simp add: sub_list_def) + +lemma sub_list_rsp2: "\xs \ ys\ \ sub_list zs xs = sub_list zs ys" + by (simp add: sub_list_def) + +lemma [quot_respect]: + "(op \ ===> op \ ===> op =) sub_list sub_list" + by (auto simp only: fun_rel_def sub_list_rsp1 sub_list_rsp2) + section {* Cardinality of finite sets *} fun