diff -r 6088fea1c8b1 -r 8a1c8dc72b5c FSet.thy --- a/FSet.thy Mon Dec 07 14:00:36 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,440 +0,0 @@ -theory FSet -imports QuotMain -begin - -inductive - list_eq (infix "\" 50) -where - "a#b#xs \ b#a#xs" -| "[] \ []" -| "xs \ ys \ ys \ xs" -| "a#a#xs \ a#xs" -| "xs \ ys \ a#xs \ a#ys" -| "\xs1 \ xs2; xs2 \ xs3\ \ xs1 \ xs3" - -lemma list_eq_refl: - shows "xs \ xs" - by (induct xs) (auto intro: list_eq.intros) - -lemma equivp_list_eq: - shows "equivp list_eq" - unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def - apply(auto intro: list_eq.intros list_eq_refl) - done - -quotient fset = "'a list" / "list_eq" - apply(rule equivp_list_eq) - done - -print_theorems - -typ "'a fset" -thm "Rep_fset" -thm "ABS_fset_def" - -quotient_def - EMPTY :: "'a fset" -where - "EMPTY \ ([]::'a list)" - -term Nil -term EMPTY -thm EMPTY_def - -quotient_def - INSERT :: "'a \ 'a fset \ 'a fset" -where - "INSERT \ op #" - -term Cons -term INSERT -thm INSERT_def - -quotient_def - FUNION :: "'a fset \ 'a fset \ 'a fset" -where - "FUNION \ (op @)" - -term append -term FUNION -thm FUNION_def - -thm Quotient_fset - -thm QUOT_TYPE_I_fset.thm11 - - -fun - membship :: "'a \ 'a list \ bool" (infix "memb" 100) -where - m1: "(x memb []) = False" -| m2: "(x memb (y#xs)) = ((x=y) \ (x memb xs))" - -fun - card1 :: "'a list \ nat" -where - card1_nil: "(card1 []) = 0" -| card1_cons: "(card1 (x # xs)) = (if (x memb xs) then (card1 xs) else (Suc (card1 xs)))" - -quotient_def - CARD :: "'a fset \ nat" -where - "CARD \ card1" - -term card1 -term CARD -thm CARD_def - -(* text {* - Maybe make_const_def should require a theorem that says that the particular lifted function - respects the relation. With it such a definition would be impossible: - make_const_def @{binding CARD} @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd -*}*) - -lemma card1_0: - fixes a :: "'a list" - shows "(card1 a = 0) = (a = [])" - by (induct a) auto - -lemma not_mem_card1: - fixes x :: "'a" - fixes xs :: "'a list" - shows "(~(x memb xs)) = (card1 (x # xs) = Suc (card1 xs))" - by auto - -lemma mem_cons: - fixes x :: "'a" - fixes xs :: "'a list" - assumes a : "x memb xs" - shows "x # xs \ xs" - using a by (induct xs) (auto intro: list_eq.intros ) - -lemma card1_suc: - fixes xs :: "'a list" - fixes n :: "nat" - assumes c: "card1 xs = Suc n" - shows "\a ys. ~(a memb ys) \ xs \ (a # ys)" - using c -apply(induct xs) -apply (metis Suc_neq_Zero card1_0) -apply (metis QUOT_TYPE_I_fset.R_trans card1_cons list_eq_refl mem_cons) -done - -definition - rsp_fold -where - "rsp_fold f = ((!u v. (f u v = f v u)) \ (!u v w. ((f u (f v w) = f (f u v) w))))" - -primrec - fold1 -where - "fold1 f (g :: 'a \ 'b) (z :: 'b) [] = z" -| "fold1 f g z (a # A) = - (if rsp_fold f - then ( - if (a memb A) then (fold1 f g z A) else (f (g a) (fold1 f g z A)) - ) else z)" - -lemma fs1_strong_cases: - fixes X :: "'a list" - shows "(X = []) \ (\a. \ Y. (~(a memb Y) \ (X \ a # Y)))" - apply (induct X) - apply (simp) - apply (metis QUOT_TYPE_I_fset.thm11 list_eq_refl mem_cons m1) - done - -quotient_def - IN :: "'a \ 'a fset \ bool" -where - "IN \ membship" - -term membship -term IN -thm IN_def - -term fold1 -quotient_def - FOLD :: "('a \ 'a \ 'a) \ ('b \ 'a) \ 'a \ 'b fset \ 'a" -where - "FOLD \ fold1" - -term fold1 -term fold -thm fold_def - -quotient_def - fmap::"('a \ 'b) \ 'a fset \ 'b fset" -where - "fmap \ map" - -term map -term fmap -thm fmap_def - -lemma memb_rsp: - fixes z - assumes a: "x \ y" - shows "(z memb x) = (z memb y)" - using a by induct auto - -lemma ho_memb_rsp[quotient_rsp]: - "(op = ===> (op \ ===> op =)) (op memb) (op memb)" - by (simp add: memb_rsp) - -lemma card1_rsp: - fixes a b :: "'a list" - assumes e: "a \ b" - shows "card1 a = card1 b" - using e by induct (simp_all add:memb_rsp) - -lemma ho_card1_rsp[quotient_rsp]: - "(op \ ===> op =) card1 card1" - by (simp add: card1_rsp) - -lemma cons_rsp[quotient_rsp]: - fixes z - assumes a: "xs \ ys" - shows "(z # xs) \ (z # ys)" - using a by (rule list_eq.intros(5)) - -lemma ho_cons_rsp[quotient_rsp]: - "(op = ===> op \ ===> op \) op # op #" - by (simp add: cons_rsp) - -lemma append_rsp_fst: - assumes a : "l1 \ l2" - shows "(l1 @ s) \ (l2 @ s)" - using a - by (induct) (auto intro: list_eq.intros list_eq_refl) - -lemma append_end: - shows "(e # l) \ (l @ [e])" - apply (induct l) - apply (auto intro: list_eq.intros list_eq_refl) - done - -lemma rev_rsp: - shows "a \ rev a" - apply (induct a) - apply simp - apply (rule list_eq_refl) - apply simp_all - apply (rule list_eq.intros(6)) - prefer 2 - apply (rule append_rsp_fst) - apply assumption - apply (rule append_end) - done - -lemma append_sym_rsp: - shows "(a @ b) \ (b @ a)" - apply (rule list_eq.intros(6)) - apply (rule append_rsp_fst) - apply (rule rev_rsp) - apply (rule list_eq.intros(6)) - apply (rule rev_rsp) - apply (simp) - apply (rule append_rsp_fst) - apply (rule list_eq.intros(3)) - apply (rule rev_rsp) - done - -lemma append_rsp: - assumes a : "l1 \ r1" - assumes b : "l2 \ r2 " - shows "(l1 @ l2) \ (r1 @ r2)" - apply (rule list_eq.intros(6)) - apply (rule append_rsp_fst) - using a apply (assumption) - apply (rule list_eq.intros(6)) - apply (rule append_sym_rsp) - apply (rule list_eq.intros(6)) - apply (rule append_rsp_fst) - using b apply (assumption) - apply (rule append_sym_rsp) - done - -lemma ho_append_rsp[quotient_rsp]: - "(op \ ===> op \ ===> op \) op @ op @" - by (simp add: append_rsp) - -lemma map_rsp: - assumes a: "a \ b" - shows "map f a \ map f b" - using a - apply (induct) - apply(auto intro: list_eq.intros) - done - -lemma ho_map_rsp[quotient_rsp]: - "(op = ===> op \ ===> op \) map map" - by (simp add: map_rsp) - -lemma map_append: - "(map f (a @ b)) \ (map f a) @ (map f b)" - by simp (rule list_eq_refl) - -lemma ho_fold_rsp[quotient_rsp]: - "(op = ===> op = ===> op = ===> op \ ===> op =) fold1 fold1" - apply (auto) - apply (case_tac "rsp_fold x") - prefer 2 - apply (erule_tac list_eq.induct) - apply (simp_all) - apply (erule_tac list_eq.induct) - apply (simp_all) - apply (auto simp add: memb_rsp rsp_fold_def) -done - -lemma list_equiv_rsp[quotient_rsp]: - shows "(op \ ===> op \ ===> op =) op \ op \" -by (auto intro: list_eq.intros) - -print_quotients - -ML {* val qty = @{typ "'a fset"} *} -ML {* val rsp_thms = - @{thms ho_memb_rsp ho_cons_rsp ho_card1_rsp ho_map_rsp ho_append_rsp ho_fold_rsp} *} - -ML {* val (rty, rel, rel_refl, rel_eqv) = lookup_quot_data @{context} qty *} -ML {* val (trans2, reps_same, absrep, quot) = lookup_quot_thms @{context} "fset"; *} -ML {* fun lift_tac_fset lthy t = lift_tac lthy t *} - -lemma "IN x EMPTY = False" -apply(tactic {* procedure_tac @{context} @{thm m1} 1 *}) -apply(tactic {* regularize_tac @{context} 1 *}) -apply(tactic {* all_inj_repabs_tac @{context} [rel_refl] [trans2] 1 *}) -apply(tactic {* clean_tac @{context} 1*}) -done - -lemma "IN x (INSERT y xa) = (x = y \ IN x xa)" -by (tactic {* lift_tac_fset @{context} @{thm m2} 1 *}) - -lemma "INSERT a (INSERT a x) = INSERT a x" -apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(4)} 1 *}) -done - -lemma "x = xa \ INSERT a x = INSERT a xa" -apply (tactic {* lift_tac_fset @{context} @{thm list_eq.intros(5)} 1 *}) -done - -lemma "CARD x = Suc n \ (\a b. \ IN a b & x = INSERT a b)" -apply (tactic {* lift_tac_fset @{context} @{thm card1_suc} 1 *}) -done - -lemma "(\ IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))" -apply (tactic {* lift_tac_fset @{context} @{thm not_mem_card1} 1 *}) -done - -lemma "FOLD f g (z::'b) (INSERT a x) = - (if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)" -apply(tactic {* lift_tac_fset @{context} @{thm fold1.simps(2)} 1 *}) -done - -ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy [rel_refl] [trans2] *} - -lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)" -apply (tactic {* lift_tac_fset @{context} @{thm map_append} 1 *}) -done - -lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)" -apply (tactic {* lift_tac_fset @{context} @{thm append_assoc} 1 *}) -done - - -lemma "\P EMPTY; \a x. P x \ P (INSERT a x)\ \ P l" -apply (tactic {* (ObjectLogic.full_atomize_tac THEN' gen_frees_tac @{context}) 1 *}) -apply(tactic {* procedure_tac @{context} @{thm list.induct} 1 *}) -apply(tactic {* regularize_tac @{context} 1 *}) -defer -apply(tactic {* clean_tac @{context} 1 *}) -apply(tactic {* inj_repabs_tac_fset @{context} 1*})+ -done - -lemma list_induct_part: - assumes a: "P (x :: 'a list) ([] :: 'c list)" - assumes b: "\e t. P x t \ P x (e # t)" - shows "P x l" - apply (rule_tac P="P x" in list.induct) - apply (rule a) - apply (rule b) - apply (assumption) - done - -ML {* quot *} -thm quotient_thm - -lemma "P (x :: 'a list) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" -apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *}) -done - -lemma "P (x :: 'a fset) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" -apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *}) -done - -lemma "P (x :: 'a fset) ([] :: 'c list) \ (\e t. P x t \ P x (e # t)) \ P x l" -apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *}) -done - -quotient fset2 = "'a list" / "list_eq" - apply(rule equivp_list_eq) - done - -quotient_def - EMPTY2 :: "'a fset2" -where - "EMPTY2 \ ([]::'a list)" - -quotient_def - INSERT2 :: "'a \ 'a fset2 \ 'a fset2" -where - "INSERT2 \ op #" - -ML {* val quot = @{thms Quotient_fset Quotient_fset2} *} -ML {* fun inj_repabs_tac_fset lthy = inj_repabs_tac lthy [rel_refl] [trans2] *} -ML {* fun lift_tac_fset lthy t = lift_tac lthy t *} - -lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" -apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *}) -done - -lemma "P (x :: 'a fset) (EMPTY2 :: 'c fset2) \ (\e t. P x t \ P x (INSERT2 e t)) \ P x l" -apply (tactic {* lift_tac_fset @{context} @{thm list_induct_part} 1 *}) -done - -quotient_def - fset_rec::"'a \ ('b \ 'b fset \ 'a \ 'a) \ 'b fset \ 'a" -where - "fset_rec \ list_rec" - -quotient_def - fset_case::"'a \ ('b \ 'b fset \ 'a) \ 'b fset \ 'a" -where - "fset_case \ list_case" - -(* Probably not true without additional assumptions about the function *) -lemma list_rec_rsp[quotient_rsp]: - "(op = ===> (op = ===> op \ ===> op =) ===> op \ ===> op =) list_rec list_rec" - apply (auto) - apply (erule_tac list_eq.induct) - apply (simp_all) - sorry - -lemma list_case_rsp[quotient_rsp]: - "(op = ===> (op = ===> op \ ===> op =) ===> op \ ===> op =) list_case list_case" - apply (auto) - sorry - -ML {* val rsp_thms = @{thms list_rec_rsp list_case_rsp} @ rsp_thms *} -ML {* fun lift_tac_fset lthy t = lift_tac lthy t *} - -lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)" -apply (tactic {* lift_tac_fset @{context} @{thm list.recs(2)} 1 *}) -done - -lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa" -apply (tactic {* lift_tac_fset @{context} @{thm list.cases(2)} 1 *}) -done - - -end