diff -r db158e995bfc -r 9df6144e281b Attic/Quot/Examples/FSet.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Attic/Quot/Examples/FSet.thy Thu Feb 25 07:57:17 2010 +0100 @@ -0,0 +1,433 @@ +theory FSet +imports "../Quotient" "../Quotient_List" "../Quotient_Product" List +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_type + 'a fset = "'a list" / "list_eq" + by (rule equivp_list_eq) + +quotient_definition + "EMPTY :: 'a fset" +is + "[]::'a list" + +quotient_definition + "INSERT :: 'a \ 'a fset \ 'a fset" +is + "op #" + +quotient_definition + "FUNION :: 'a fset \ 'a fset \ 'a fset" +is + "op @" + +fun + card1 :: "'a list \ nat" +where + card1_nil: "(card1 []) = 0" +| card1_cons: "(card1 (x # xs)) = (if (x mem xs) then (card1 xs) else (Suc (card1 xs)))" + +quotient_definition + "CARD :: 'a fset \ nat" +is + "card1" + +quotient_definition + "fconcat :: ('a fset) fset \ 'a fset" +is + "concat" + +term concat +term fconcat + +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 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 mem xs)) = (card1 (x # xs) = Suc (card1 xs))" + by auto + +lemma mem_cons: + fixes x :: "'a" + fixes xs :: "'a list" + assumes a : "x mem 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 mem ys) \ xs \ (a # ys)" + using c +apply(induct xs) +apply (metis Suc_neq_Zero card1_0) +apply (metis FSet.card1_cons list_eq.intros(6) 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 mem 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 mem Y) \ (X \ a # Y)))" + apply (induct X) + apply (simp) + apply (metis List.member.simps(1) list_eq.intros(6) list_eq_refl mem_cons) + done + +quotient_definition + "IN :: 'a \ 'a fset \ bool" +is + "op mem" + +quotient_definition + "FOLD :: ('a \ 'a \ 'a) \ ('b \ 'a) \ 'a \ 'b fset \ 'a" +is + "fold1" + +quotient_definition + "fmap :: ('a \ 'b) \ 'a fset \ 'b fset" +is + "map" + +lemma mem_rsp: + fixes z + assumes a: "x \ y" + shows "(z mem x) = (z mem y)" + using a by induct auto + +lemma ho_memb_rsp[quot_respect]: + "(op = ===> (op \ ===> op =)) (op mem) (op mem)" + by (simp add: mem_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: mem_rsp) + +lemma ho_card1_rsp[quot_respect]: + "(op \ ===> op =) card1 card1" + by (simp add: card1_rsp) + +lemma cons_rsp: + fixes z + assumes a: "xs \ ys" + shows "(z # xs) \ (z # ys)" + using a by (rule list_eq.intros(5)) + +lemma ho_cons_rsp[quot_respect]: + "(op = ===> op \ ===> op \) op # op #" + by (simp add: cons_rsp) + +lemma append_rsp_aux1: + assumes a : "l2 \ r2 " + shows "(h @ l2) \ (h @ r2)" +using a +apply(induct h) +apply(auto intro: list_eq.intros(5)) +done + +lemma append_rsp_aux2: + assumes a : "l1 \ r1" "l2 \ r2 " + shows "(l1 @ l2) \ (r1 @ r2)" +using a +apply(induct arbitrary: l2 r2) +apply(simp_all) +apply(blast intro: list_eq.intros append_rsp_aux1)+ +done + +lemma append_rsp[quot_respect]: + "(op \ ===> op \ ===> op \) op @ op @" + by (auto simp add: append_rsp_aux2) + +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[quot_respect]: + "(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[quot_respect]: + "(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: mem_rsp rsp_fold_def) +done + +lemma list_equiv_rsp[quot_respect]: + shows "(op \ ===> op \ ===> op =) op \ op \" +by (auto intro: list_eq.intros) + +lemma "IN x EMPTY = False" +apply(lifting member.simps(1)) +done + +lemma "IN x (INSERT y xa) = (x = y \ IN x xa)" +apply (lifting member.simps(2)) +done + +lemma "INSERT a (INSERT a x) = INSERT a x" +apply (lifting list_eq.intros(4)) +done + +lemma "x = xa \ INSERT a x = INSERT a xa" +apply (lifting list_eq.intros(5)) +done + +lemma "CARD x = Suc n \ (\a b. \ IN a b & x = INSERT a b)" +apply (lifting card1_suc) +done + +lemma "(\ IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))" +apply (lifting not_mem_card1) +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(lifting fold1.simps(2)) +done + +lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)" +apply (lifting map_append) +done + +lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)" +apply (lifting append_assoc) +done + + +lemma "\P EMPTY; \a x. P x \ P (INSERT a x)\ \ P l" +apply(lifting list.induct) +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 + +lemma "P (x :: 'a list) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" +apply (lifting list_induct_part) +done + +lemma "P (x :: 'a fset) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" +apply (lifting list_induct_part) +done + +lemma "P (x :: 'a fset) ([] :: 'c list) \ (\e t. P x t \ P x (e # t)) \ P x l" +apply (lifting list_induct_part) +done + +quotient_type 'a fset2 = "'a list" / "list_eq" + by (rule equivp_list_eq) + +quotient_definition + "EMPTY2 :: 'a fset2" +is + "[]::'a list" + +quotient_definition + "INSERT2 :: 'a \ 'a fset2 \ 'a fset2" +is + "op #" + +lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \ (\e t. P x t \ P x (INSERT e t)) \ P x l" +apply (lifting list_induct_part) +done + +lemma "P (x :: 'a fset) (EMPTY2 :: 'c fset2) \ (\e t. P x t \ P x (INSERT2 e t)) \ P x l" +apply (lifting list_induct_part) +done + +quotient_definition + "fset_rec :: 'a \ ('b \ 'b fset \ 'a \ 'a) \ 'b fset \ 'a" +is + "list_rec" + +quotient_definition + "fset_case :: 'a \ ('b \ 'b fset \ 'a) \ 'b fset \ 'a" +is + "list_case" + +(* Probably not true without additional assumptions about the function *) +lemma list_rec_rsp[quot_respect]: + "(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[quot_respect]: + "(op = ===> (op = ===> op \ ===> op =) ===> op \ ===> op =) list_case list_case" + apply (auto) + sorry + +lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)" +apply (lifting list.recs(2)) +done + +lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa" +apply (lifting list.cases(2)) +done + +lemma ttt: "((op @) x ((op #) e [])) = (((op #) e x))" +sorry + +lemma "(FUNION x (INSERT e EMPTY)) = ((INSERT e x))" +apply (lifting ttt) +done + + +lemma ttt2: "(\e. ((op @) x ((op #) e []))) = (\e. ((op #) e x))" +sorry + +lemma "(\e. (FUNION x (INSERT e EMPTY))) = (\e. (INSERT e x))" +apply(lifting ttt2) +apply(regularize) +apply(rule impI) +apply(simp) +apply(rule allI) +apply(rule list_eq_refl) +done + +lemma ttt3: "(\x. ((op @) x (e # []))) = (op #) e" +sorry + +lemma "(\x. (FUNION x (INSERT e EMPTY))) = INSERT e" +apply(lifting ttt3) +apply(regularize) +apply(auto simp add: cons_rsp) +done +lemma hard: "(\P. \Q. P (Q (x::'a list))) = (\P. \Q. Q (P (x::'a list)))" +sorry + +lemma eq_imp_rel: + shows "equivp R \ a = b \ R a b" + by (simp add: equivp_reflp) + + +lemma hard_lift: "(\P. \Q. P (Q (x::'a fset))) = (\P. \Q. Q (P (x::'a fset)))" +apply(lifting hard) +apply(regularize) +apply(rule fun_rel_id_asm) +apply(subst babs_simp) +apply(tactic {* Quotient_Tacs.quotient_tac @{context} 1 *}) +apply(rule fun_rel_id_asm) +apply(rule impI) +apply(rule mp[OF eq_imp_rel[OF fset_equivp]]) +apply(drule fun_cong) +apply(drule fun_cong) +apply(assumption) +done + +lemma test: "All (\(x::'a list, y). x = y)" +sorry + +lemma "All (\(x::'a fset, y). x = y)" +apply(lifting test) +done + +lemma test2: "Ex (\(x::'a list, y). x = y)" +sorry + +lemma "Ex (\(x::'a fset, y). x = y)" +apply(lifting test2) +done + +lemma test3: "All (\ (x :: 'a list, y, z). x = y \ y = z)" +sorry + +lemma "All (\ (x :: 'a fset, y, z). x = y \ y = z)" +apply(lifting test3) +done + +lemma test4: "\ (x :: 'a list, y, z) \ Respects( + prod_rel (op \) (prod_rel (op \) (op \)) +). x = y \ y = z" +sorry + +lemma "All (\ (x :: 'a fset, y, z). x = y \ y = z)" +apply (lifting test4) +sorry + +lemma test5: "\ (x :: 'a list \ 'a list, y) \ Respects( + prod_rel (op \ ===> op \) (op \ ===> op \) +). (op \ ===> op \) x y" +sorry + +lemma "All (\ (x :: 'a fset \ 'a fset, y). x = y)" +apply (lifting test5) +done + +lemma test6: "\ (x :: 'a list \ 'a list, y, z) \ Respects( + prod_rel (op \ ===> op \) (prod_rel (op \ ===> op \) (op \ ===> op \)) +). (op \ ===> op \) x y \ (op \ ===> op \) y z" +sorry + +lemma "All (\ (x :: 'a fset \ 'a fset, y, z). x = y \ y = z)" +apply (lifting test6) +done + +end