diff -r 0d9e8aa1bc7a -r 88094aa77026 Quot/Examples/FSet2.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Quot/Examples/FSet2.thy Thu Dec 10 03:48:39 2009 +0100 @@ -0,0 +1,119 @@ +theory FSet2 +imports "../QuotMain" 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 +by (auto intro: list_eq.intros list_eq_refl) + +quotient fset = "'a list" / "list_eq" + by (rule equivp_list_eq) + +quotient_def + fempty :: "fempty :: 'a fset" ("{||}") +where + "[]" + +quotient_def + finsert :: "finsert :: 'a \ 'a fset \ 'a fset" +where + "(op #)" + +lemma finsert_rsp[quot_respect]: + shows "(op = ===> op \ ===> op \) (op #) (op #)" +by (auto intro: list_eq.intros) + +quotient_def + funion :: "funion :: 'a fset \ 'a fset \ 'a fset" ("_ \f _" [65,66] 65) +where + "(op @)" + +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]: + shows "(op \ ===> op \ ===> op \) op @ op @" + by (auto simp add: append_rsp_aux2) + + +quotient_def + fmem :: " fmem :: 'a \ 'a fset \ bool" ("_ \f _" [50, 51] 50) +where + "(op mem)" + +lemma memb_rsp_aux: + assumes a: "x \ y" + shows "(z mem x) = (z mem y)" + using a by induct auto + +lemma memb_rsp[quot_respect]: + shows "(op = ===> (op \ ===> op =)) (op mem) (op mem)" + by (simp add: memb_rsp_aux) + +definition + fnot_mem :: "'a \ 'a fset \ bool" ("_ \f _" [50, 51] 50) +where + "x \f S \ \(x \f S)" + +definition + "inter_list" :: "'a list \ 'a list \ 'a list" +where + "inter_list X Y \ [x \ X. x\set Y]" + +quotient_def + finter :: "finter::'a fset \ 'a fset \ 'a fset" ("_ \f _" [70, 71] 70) +where + "inter_list" + +no_syntax + "@Finset" :: "args => 'a fset" ("{|(_)|}") +syntax + "@Finfset" :: "args => 'a fset" ("{|(_)|}") +translations + "{|x, xs|}" == "CONST finsert x {|xs|}" + "{|x|}" == "CONST finsert x {||}" + + +subsection {* Empty sets *} + +lemma test: + shows "\(x # xs \ [])" +sorry + +lemma finsert_not_empty[simp]: + shows "finsert x S \ {||}" + by (lifting test) + + + + +end;