# HG changeset patch # User Christian Urban # Date 1271846300 -7200 # Node ID ec45c81d1ca6102cb9d81416c26579dfca3c97ac # Parent d96c48fd731661360fb021f85545c42b50396974# Parent efbc22a6f1a40e0b640548905ace27aa82acd70c merged diff -r d96c48fd7316 -r ec45c81d1ca6 Attic/Quot/Examples/FSet3.thy --- a/Attic/Quot/Examples/FSet3.thy Wed Apr 21 12:37:44 2010 +0200 +++ b/Attic/Quot/Examples/FSet3.thy Wed Apr 21 12:38:20 2010 +0200 @@ -1,291 +1,73 @@ 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 +lemma list_rel_find_element: + assumes a: "x \ set a" + and b: "list_rel R a b" + shows "\y. (y \ set b \ R x y)" +proof - + have "length a = length b" using b by (rule list_rel_len) + then show ?thesis using a b proof (induct a b rule: list_induct2) + case Nil + show ?case using Nil.prems by simp next - case False - have "x |\| S" by fact - then show "P (finsert x S)" using prem2 asm by simp + case (Cons x xs y ys) + show ?case using Cons by auto 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) - -lemma concat2: - shows "concat (x # xs) \ x @ concat xs" -by (simp) - -lemma concat_rsp: - "\list_rel op \ x x'; x' \ y'; list_rel op \ y' y\ \ concat x \ concat y" - apply (induct x y arbitrary: x' y' rule: list_induct2') +lemma concat_rsp_pre: + "\list_rel op \ x x'; x' \ y'; list_rel op \ y' y; \x\set x. xa \ set x\ \ + \x\set y. xa \ set x" + apply clarify + apply (frule list_rel_find_element[of _ "x"]) + apply assumption + apply clarify + apply (subgoal_tac "ya \ set y'") + prefer 2 apply simp - defer defer - apply (simp only: concat.simps) - sorry + apply (frule list_rel_find_element[of _ "y'"]) + apply assumption + apply auto + done lemma [quot_respect]: shows "(list_rel op \ OOO op \ ===> op \) concat concat" apply (simp only: fun_rel_def) apply clarify - apply (rule concat_rsp) + apply (simp (no_asm)) + apply rule + apply rule + apply (erule concat_rsp_pre) apply assumption+ + apply (rule concat_rsp_pre) + prefer 4 + apply assumption + apply (rule list_rel_symp[OF list_eq_equivp]) + apply assumption + apply (rule equivp_symp[OF list_eq_equivp]) + apply assumption + apply (rule list_rel_symp[OF list_eq_equivp]) + apply assumption 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) - apply auto - done + by (rule eq_reflection) auto lemma map_rel_cong: "b \ ba \ map f b \ map f ba" unfolding list_eq.simps - apply(simp only: set_map set_in_eq) - done + by (simp only: set_map set_in_eq) lemma quotient_compose_list_pre: "(list_rel op \ OOO op \) r s = @@ -342,7 +124,7 @@ (abs_fset \ (map abs_fset)) ((map rep_fset) \ rep_fset)" unfolding Quotient_def comp_def apply (rule)+ - apply (simp add: abs_o_rep[OF Quotient_fset] id_simps Quotient_abs_rep[OF Quotient_fset]) + apply (simp add: abs_o_rep[OF Quotient_fset] Quotient_abs_rep[OF Quotient_fset] map_id) apply (rule) apply (rule) apply (rule) @@ -359,11 +141,9 @@ lemma fconcat_empty: shows "fconcat {||} = {||}" - apply(lifting concat1) + apply(lifting concat.simps(1)) 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 +153,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]: @@ -384,30 +162,12 @@ lemma fconcat_insert: shows "fconcat (finsert x S) = x |\| fconcat S" - apply(lifting concat2) - apply(cleaning) + apply (lifting concat.simps(2)) + apply (cleaning) apply (simp add: finsert_def fconcat_def comp_def) - apply cleaning + apply (cleaning) done -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) - - -text {* lifted section *} - (* TBD *) text {* syntax for fset comprehensions (adapted from lists) *} @@ -507,210 +267,4 @@ (* BELOW CONSTRUCTION SITE *) -lemma no_mem_nil: - "(\a. a \ set A) = (A = [])" -by (induct A) (auto) - -lemma none_mem_nil: - "(\a. a \ set A) = (A \ [])" -by simp - -lemma mem_cons: - "a \ set A \ a # A \ A" -by auto - -lemma cons_left_comm: - "x # y # A \ y # x # A" -by (auto simp add: id_simps) - -lemma cons_left_idem: - "x # x # A \ x # A" -by (auto simp add: id_simps) - -lemma finite_set_raw_strong_cases: - "(X = []) \ (\a Y. ((a \ set Y) \ (X \ a # Y)))" - apply (induct X) - apply (simp) - apply (rule disjI2) - apply (erule disjE) - apply (rule_tac x="a" in exI) - apply (rule_tac x="[]" in exI) - apply (simp) - apply (erule exE)+ - apply (case_tac "a = aa") - apply (rule_tac x="a" in exI) - apply (rule_tac x="Y" in exI) - apply (simp) - apply (rule_tac x="aa" in exI) - apply (rule_tac x="a # Y" in exI) - apply (auto) - done - -fun - delete_raw :: "'a list \ 'a \ 'a list" -where - "delete_raw [] x = []" -| "delete_raw (a # A) x = (if (a = x) then delete_raw A x else a # (delete_raw A x))" - -lemma mem_delete_raw: - "x \ set (delete_raw A a) = (x \ set A \ \(x = a))" - by (induct A arbitrary: x a) (auto) - -lemma mem_delete_raw_ident: - "\(a \ set (delete_raw A a))" -by (induct A) (auto) - -lemma not_mem_delete_raw_ident: - "b \ set A \ (delete_raw A b = A)" -by (induct A) (auto) - -lemma delete_raw_RSP: - "A \ B \ delete_raw A a \ delete_raw B a" -apply(induct A arbitrary: B a) -apply(auto) -sorry - -lemma cons_delete_raw: - "a # (delete_raw A a) \ (if a \ set A then A else (a # A))" -sorry - -lemma mem_cons_delete_raw: - "a \ set A \ a # (delete_raw A a) \ A" -sorry - -lemma finite_set_raw_delete_raw_cases: - "X = [] \ (\a. a mem X \ X \ a # delete_raw X a)" - by (induct X) (auto) - - - - - -lemma list2set_thm: - shows "set [] = {}" - and "set (h # t) = insert h (set t)" - by (auto) - -lemma list2set_RSP: - "A \ B \ set A = set B" - by auto - -definition - rsp_fold -where - "rsp_fold f = (\u v w. (f u (f v w) = f v (f u w)))" - -primrec - fold_raw :: "('a \ 'b \ 'b) \ 'b \ 'a list \ 'b" -where - "fold_raw f z [] = z" -| "fold_raw f z (a # A) = - (if (rsp_fold f) then - if a mem A then fold_raw f z A - else f a (fold_raw f z A) - else z)" - -lemma mem_lcommuting_fold_raw: - "rsp_fold f \ h mem B \ fold_raw f z B = f h (fold_raw f z (delete_raw B h))" -sorry - -lemma fold_rsp[quot_respect]: - "(op = ===> op = ===> op \ ===> op =) fold_raw fold_raw" -apply(auto) -sorry - -primrec - inter_raw -where - "inter_raw [] B = []" -| "inter_raw (a # A) B = (if a mem B then a # inter_raw A B else inter_raw A B)" - -lemma mem_inter_raw: - "x mem (inter_raw A B) = x mem A \ x mem B" -sorry - -lemma inter_raw_RSP: - "A1 \ A2 \ B1 \ B2 \ (inter_raw A1 B1) \ (inter_raw A2 B2)" -sorry - - -(* LIFTING DEFS *) - - -section {* Constants on the Quotient Type *} - - -quotient_definition - "fdelete :: 'a fset \ 'a \ 'a fset" - is "delete_raw" - -quotient_definition - finter ("_ \f _" [70, 71] 70) -where - "finter :: 'a fset \ 'a fset \ 'a fset" - is "inter_raw" - -quotient_definition - "ffold :: ('a \ 'b \ 'b) \ 'b \ 'a fset \ 'b" - is "fold_raw" - -quotient_definition - "fset_to_set :: 'a fset \ 'a set" - is "set" - - -section {* Lifted Theorems *} - -thm list.cases (* ??? *) - -thm cons_left_comm -lemma "finsert a (finsert b S) = finsert b (finsert a S)" -by (lifting cons_left_comm) - -thm cons_left_idem -lemma "finsert a (finsert a S) = finsert a S" -by (lifting cons_left_idem) - -(* thm MEM: - MEM x [] = F - MEM x (h::t) = (x=h) \/ MEM x t *) -thm none_mem_nil -(*lemma "(\a. a \f A) = (A = fempty)"*) - -thm mem_cons -thm finite_set_raw_strong_cases -(*thm card_raw.simps*) -(*thm not_mem_card_raw*) -(*thm card_raw_suc*) - -lemma "fcard X = Suc n \ (\a S. a \f S & X = finsert a S)" -(*by (lifting card_raw_suc)*) -sorry - -(*thm card_raw_cons_gt_0 -thm mem_card_raw_gt_0 -thm not_nil_equiv_cons -*) -thm delete_raw.simps -(*thm mem_delete_raw*) -(*thm card_raw_delete_raw*) -thm cons_delete_raw -thm mem_cons_delete_raw -thm finite_set_raw_delete_raw_cases -thm append.simps -(* MEM_APPEND: MEM e (APPEND l1 l2) = MEM e l1 \/ MEM e l2 *) -thm inter_raw.simps -thm mem_inter_raw -thm fold_raw.simps -thm list2set_thm -thm list_eq_def -thm list.induct -lemma "\P fempty; \a x. P x \ P (finsert a x)\ \ P l" -by (lifting list.induct) - -(* We also have map and some properties of it in FSet *) -(* and the following which still lifts ok *) -lemma "funion (funion x xa) xb = funion x (funion xa xb)" -by (lifting append_assoc) - end