# HG changeset patch # User Cezary Kaliszyk # Date 1271837597 -7200 # Node ID c50601bc617e2ae61aa73e01cfbd48d9817d6fc1 # Parent 9972dc5bd7ad1b0336fd7239a2af1ba1f3b61280 Remove the part already in FSet and leave the experiments diff -r 9972dc5bd7ad -r c50601bc617e Attic/Quot/Examples/FSet3.thy --- a/Attic/Quot/Examples/FSet3.thy Wed Apr 21 09:48:35 2010 +0200 +++ b/Attic/Quot/Examples/FSet3.thy Wed Apr 21 10:13:17 2010 +0200 @@ -1,257 +1,22 @@ theory FSet3 -imports "../Quotient" "../Quotient_List" List +imports "../../../Nominal/FSet" begin -ML {* -structure QuotientRules = Named_Thms - (val name = "quot_thm" - val description = "Quotient theorems.") -*} - -ML {* -open QuotientRules -*} - -fun - list_eq :: "'a list \ 'a list \ bool" (infix "\" 50) -where - "list_eq xs ys = (\x. x \ set xs \ x \ set ys)" - -lemma list_eq_equivp: - shows "equivp list_eq" -unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def -by auto - -(* FIXME-TODO: because of beta-reduction, one cannot give the *) -(* FIXME-TODO: relation as a term or abbreviation *) -quotient_type - 'a fset = "'a list" / "list_eq" -by (rule list_eq_equivp) - - -section {* empty fset, finsert and membership *} - -quotient_definition - fempty ("{||}") -where - "fempty :: 'a fset" -is "[]::'a list" - -quotient_definition - "finsert :: 'a \ 'a fset \ 'a fset" -is "op #" - -syntax - "@Finset" :: "args => 'a fset" ("{|(_)|}") - -translations - "{|x, xs|}" == "CONST finsert x {|xs|}" - "{|x|}" == "CONST finsert x {||}" - -definition - memb :: "'a \ 'a list \ bool" -where - "memb x xs \ x \ set xs" - -quotient_definition - fin ("_ |\| _" [50, 51] 50) -where - "fin :: 'a \ 'a fset \ bool" -is "memb" - -abbreviation - fnotin :: "'a \ 'a fset \ bool" ("_ |\| _" [50, 51] 50) -where - "a |\| S \ \(a |\| S)" - -lemma memb_rsp[quot_respect]: - shows "(op = ===> op \ ===> op =) memb memb" -by (auto simp add: memb_def) - -lemma nil_rsp[quot_respect]: - shows "[] \ []" -by simp - -lemma cons_rsp: - "xa \ ya \ y # xa \ y # ya" - by simp - -lemma [quot_respect]: - shows "(op = ===> op \ ===> op \) op # op #" - by simp - - -section {* Augmenting a set -- @{const finsert} *} - -text {* raw section *} - -lemma nil_not_cons: - shows "\[] \ x # xs" - by auto - -lemma memb_cons_iff: - shows "memb x (y # xs) = (x = y \ memb x xs)" - by (induct xs) (auto simp add: memb_def) - -lemma memb_consI1: - shows "memb x (x # xs)" - by (simp add: memb_def) - -lemma memb_consI2: - shows "memb x xs \ memb x (y # xs)" - by (simp add: memb_def) - -lemma memb_absorb: - shows "memb x xs \ x # xs \ xs" - by (induct xs) (auto simp add: memb_def id_simps) - -text {* lifted section *} - -lemma fin_finsert_iff[simp]: - "x |\| finsert y S = (x = y \ x |\| S)" -by (lifting memb_cons_iff) - -lemma - shows finsertI1: "x |\| finsert x S" - and finsertI2: "x |\| S \ x |\| finsert y S" - by (lifting memb_consI1, lifting memb_consI2) - - -lemma finsert_absorb [simp]: - shows "x |\| S \ finsert x S = S" - by (lifting memb_absorb) - - -section {* Singletons *} - -text {* raw section *} - -lemma singleton_list_eq: - shows "[x] \ [y] \ x = y" - by (simp add: id_simps) auto - -text {* lifted section *} - -lemma fempty_not_finsert[simp]: - shows "{||} \ finsert x S" - by (lifting nil_not_cons) - -lemma fsingleton_eq[simp]: - shows "{|x|} = {|y|} \ x = y" - by (lifting singleton_list_eq) - -section {* Union *} - -quotient_definition - funion (infixl "|\|" 65) -where - "funion :: 'a fset \ 'a fset \ 'a fset" -is - "op @" - -section {* Cardinality of finite sets *} - -fun - fcard_raw :: "'a list \ 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))" - -quotient_definition - "fcard :: 'a fset \ nat" -is "fcard_raw" - -text {* raw section *} - -lemma fcard_raw_ge_0: - assumes a: "x \ set xs" - shows "0 < fcard_raw xs" -using a -by (induct xs) (auto simp add: memb_def) - -lemma fcard_raw_delete_one: - "fcard_raw ([x \ xs. x \ y]) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)" -by (induct xs) (auto dest: fcard_raw_ge_0 simp add: memb_def) - -lemma fcard_raw_rsp_aux: - assumes a: "a \ b" - shows "fcard_raw a = fcard_raw b" -using a -apply(induct a arbitrary: b) -apply(auto simp add: memb_def) -apply(metis) -apply(drule_tac x="[x \ b. x \ a1]" in meta_spec) -apply(simp add: fcard_raw_delete_one) -apply(metis Suc_pred' fcard_raw_ge_0 fcard_raw_delete_one memb_def) -done - -lemma fcard_raw_rsp[quot_respect]: - "(op \ ===> op =) fcard_raw fcard_raw" - by (simp add: fcard_raw_rsp_aux) - -text {* lifted section *} - -lemma fcard_fempty [simp]: - shows "fcard {||} = 0" -by (lifting fcard_raw_nil) - -lemma fcard_finsert_if [simp]: - shows "fcard (finsert x S) = (if x |\| S then fcard S else Suc (fcard S))" -by (lifting fcard_raw_cons) - - -section {* Induction and Cases rules for finite sets *} +notation + list_eq (infix "\" 50) lemma fset_exhaust[case_names fempty finsert, cases type: fset]: shows "\S = {||} \ P; \x S'. S = finsert x S' \ P\ \ P" by (lifting list.exhaust) -lemma fset_induct[case_names fempty finsert]: - shows "\P {||}; \x S. P S \ P (finsert x S)\ \ P S" -by (lifting list.induct) - -lemma fset_induct2[case_names fempty finsert, induct type: fset]: - assumes prem1: "P {||}" - and prem2: "\x S. \x |\| S; P S\ \ P (finsert x S)" - shows "P S" -proof(induct S rule: fset_induct) - case fempty - show "P {||}" by (rule prem1) -next - case (finsert x S) - have asm: "P S" by fact - show "P (finsert x S)" - proof(cases "x |\| S") - case True - have "x |\| S" by fact - then show "P (finsert x S)" using asm by simp - next - case False - have "x |\| S" by fact - then show "P (finsert x S)" using prem2 asm by simp - qed -qed - - -section {* fmap and fset comprehension *} - -quotient_definition - "fmap :: ('a \ 'b) \ 'a fset \ 'b fset" -is - "map" - -quotient_definition - "fconcat :: ('a fset) fset => 'a fset" -is - "concat" - (* PROBLEM: these lemmas needs to be restated, since *) (* concat.simps(1) and concat.simps(2) contain the *) (* type variables ?'a1.0 (which are turned into frees *) (* 'a_1 *) + lemma concat1: shows "concat [] \ []" -by (simp add: id_simps) +by (simp) lemma concat2: shows "concat (x # xs) \ x @ concat xs" @@ -274,8 +39,7 @@ done lemma nil_rsp2[quot_respect]: "(list_rel op \ OOO op \) [] []" - apply (metis FSet3.nil_rsp list_rel.simps(1) pred_comp.intros) - done + by (metis nil_rsp list_rel.simps(1) pred_compI) lemma set_in_eq: "(\e. ((e \ A) \ (e \ B))) \ A = B" apply (rule eq_reflection) @@ -361,9 +125,7 @@ shows "fconcat {||} = {||}" apply(lifting concat1) apply(cleaning) - apply(simp add: comp_def) - apply(fold fempty_def[simplified id_simps]) - apply(rule refl) + apply(simp add: comp_def bot_fset_def) done lemma insert_rsp2[quot_respect]: @@ -373,9 +135,7 @@ apply (rule_tac b="x # b" in pred_compI) apply auto apply (rule_tac b="x # ba" in pred_compI) - apply (rule cons_rsp) - apply simp - apply (auto)[1] + apply auto done lemma append_rsp[quot_respect]: