# HG changeset patch # User Christian Urban # Date 1264499624 -3600 # Node ID da5e4b8317c77985fe5929a6c0e5cca9e1c4c1b8 # Parent c96e007b512ffb7bf0530dd25bbe613a91afa0ba# Parent 0879d144aaa3961fda32ba2db8def0282e341223 merged diff -r 0879d144aaa3 -r da5e4b8317c7 Quot/QuotBase.thy --- a/Quot/QuotBase.thy Tue Jan 26 09:54:43 2010 +0100 +++ b/Quot/QuotBase.thy Tue Jan 26 10:53:44 2010 +0100 @@ -72,10 +72,9 @@ where "r1 OOO r2 \ r1 OO r2 OO r1" -lemma eq_comp_r: "(op = OO R OO op =) = R" - apply (rule ext)+ - apply auto - done +lemma eq_comp_r: + shows "((op =) OOO R) = R" + by (auto simp add: expand_fun_eq) section {* Respects predicate *} diff -r 0879d144aaa3 -r da5e4b8317c7 Quot/QuotList.thy --- a/Quot/QuotList.thy Tue Jan 26 09:54:43 2010 +0100 +++ b/Quot/QuotList.thy Tue Jan 26 10:53:44 2010 +0100 @@ -1,4 +1,4 @@ -theory QuotList + theory QuotList imports QuotMain List begin @@ -12,23 +12,60 @@ declare [[map list = (map, list_rel)]] + + +text {* should probably be in Sum_Type.thy *} +lemma split_list_all: + shows "(\x. P x) \ P [] \ (\x xs. P (x#xs))" + apply(auto) + apply(case_tac x) + apply(simp_all) + done + +lemma map_id[id_simps]: + shows "map id = id" + apply(simp add: expand_fun_eq) + apply(rule allI) + apply(induct_tac x) + apply(simp_all) + done + + +lemma list_rel_reflp: + shows "equivp R \ list_rel R xs xs" + apply(induct xs) + apply(simp_all add: equivp_reflp) + done + +lemma list_rel_symp: + assumes a: "equivp R" + shows "list_rel R xs ys \ list_rel R ys xs" + apply(induct xs ys rule: list_induct2') + apply(simp_all) + apply(rule equivp_symp[OF a]) + apply(simp) + done + +lemma list_rel_transp: + assumes a: "equivp R" + shows "list_rel R xs1 xs2 \ list_rel R xs2 xs3 \ list_rel R xs1 xs3" + apply(induct xs1 xs2 arbitrary: xs3 rule: list_induct2') + apply(simp_all) + apply(case_tac xs3) + apply(simp_all) + apply(rule equivp_transp[OF a]) + apply(auto) + done + lemma list_equivp[quot_equiv]: assumes a: "equivp R" shows "equivp (list_rel R)" - unfolding equivp_def - apply(rule allI)+ - apply(induct_tac x y rule: list_induct2') - apply(simp_all add: expand_fun_eq) - apply(metis list_rel.simps(1) list_rel.simps(2)) - apply(metis list_rel.simps(1) list_rel.simps(2)) - apply(rule iffI) - apply(rule allI) - apply(case_tac x) - apply(simp_all) - using a - apply(unfold equivp_def) - apply(auto)[1] - apply(metis list_rel.simps(4)) + apply(rule equivpI) + unfolding reflp_def symp_def transp_def + apply(subst split_list_all) + apply(simp add: equivp_reflp[OF a] list_rel_reflp[OF a]) + apply(blast intro: list_rel_symp[OF a]) + apply(blast intro: list_rel_transp[OF a]) done lemma list_rel_rel: @@ -44,11 +81,8 @@ assumes q: "Quotient R Abs Rep" shows "Quotient (list_rel R) (map Abs) (map Rep)" unfolding Quotient_def - apply(rule conjI) - apply(rule allI) - apply(induct_tac a) - apply(simp) - apply(simp add: Quotient_abs_rep[OF q]) + apply(subst split_list_all) + apply(simp add: Quotient_abs_rep[OF q] abs_o_rep[OF q] map_id) apply(rule conjI) apply(rule allI) apply(induct_tac a) @@ -59,11 +93,6 @@ apply(rule list_rel_rel[OF q]) done -lemma map_id[id_simps]: "map id = id" - apply (rule ext) - apply (rule_tac list="x" in list.induct) - apply (simp_all) - done lemma cons_prs_aux: assumes q: "Quotient R Abs Rep" @@ -73,12 +102,13 @@ lemma cons_prs[quot_preserve]: assumes q: "Quotient R Abs Rep" shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)" - by (simp only: expand_fun_eq fun_map_def cons_prs_aux[OF q]) (simp) + by (simp only: expand_fun_eq fun_map_def cons_prs_aux[OF q]) + (simp) lemma cons_rsp[quot_respect]: assumes q: "Quotient R Abs Rep" - shows "(R ===> list_rel R ===> list_rel R) op # op #" - by auto + shows "(R ===> list_rel R ===> list_rel R) (op #) (op #)" + by (auto) lemma nil_prs[quot_preserve]: assumes q: "Quotient R Abs Rep" @@ -94,14 +124,16 @@ assumes a: "Quotient R1 abs1 rep1" and b: "Quotient R2 abs2 rep2" shows "(map abs2) (map ((abs1 ---> rep2) f) (map rep1 l)) = map f l" - by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) + by (induct l) + (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) lemma map_prs[quot_preserve]: assumes a: "Quotient R1 abs1 rep1" and b: "Quotient R2 abs2 rep2" shows "((abs1 ---> rep2) ---> (map rep1) ---> (map abs2)) map = map" - by (simp only: expand_fun_eq fun_map_def map_prs_aux[OF a b]) (simp) + by (simp only: expand_fun_eq fun_map_def map_prs_aux[OF a b]) + (simp) lemma map_rsp[quot_respect]: @@ -126,7 +158,8 @@ assumes a: "Quotient R1 abs1 rep1" and b: "Quotient R2 abs2 rep2" shows "((abs1 ---> abs2 ---> rep2) ---> (map rep1) ---> rep2 ---> abs2) foldr = foldr" - by (simp only: expand_fun_eq fun_map_def foldr_prs_aux[OF a b]) (simp) + by (simp only: expand_fun_eq fun_map_def foldr_prs_aux[OF a b]) + (simp) lemma foldl_prs_aux: assumes a: "Quotient R1 abs1 rep1" @@ -134,18 +167,20 @@ shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l" by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) + lemma foldl_prs[quot_preserve]: assumes a: "Quotient R1 abs1 rep1" and b: "Quotient R2 abs2 rep2" shows "((abs1 ---> abs2 ---> rep1) ---> rep1 ---> (map rep2) ---> abs1) foldl = foldl" - by (simp only: expand_fun_eq fun_map_def foldl_prs_aux[OF a b]) (simp) + by (simp only: expand_fun_eq fun_map_def foldl_prs_aux[OF a b]) + (simp) -lemma list_rel_empty: - "list_rel R [] b \ length b = 0" +lemma list_rel_empty: + shows "list_rel R [] b \ length b = 0" by (induct b) (simp_all) -lemma list_rel_len: - "list_rel R a b \ length a = length b" +lemma list_rel_len: + shows "list_rel R a b \ length a = length b" apply (induct a arbitrary: b) apply (simp add: list_rel_empty) apply (case_tac b) diff -r 0879d144aaa3 -r da5e4b8317c7 Quot/QuotOption.thy --- a/Quot/QuotOption.thy Tue Jan 26 09:54:43 2010 +0100 +++ b/Quot/QuotOption.thy Tue Jan 26 10:53:44 2010 +0100 @@ -1,7 +1,12 @@ +(* Title: QuotOption.thy + Author: Cezary Kaliszyk and Christian Urban +*) theory QuotOption imports QuotMain begin +section {* Quotient infrastructure for option type *} + fun option_rel where @@ -10,28 +15,25 @@ | "option_rel R None (Some x) = False" | "option_rel R (Some x) (Some y) = R x y" -fun - option_map -where - "option_map f None = None" -| "option_map f (Some x) = Some (f x)" +declare [[map option = (Option.map, option_rel)]] -declare [[map option = (option_map, option_rel)]] - +text {* should probably be in Option.thy *} +lemma split_option_all: + shows "(\x. P x) \ P None \ (\a. P (Some a))" +apply(auto) +apply(case_tac x) +apply(simp_all) +done lemma option_quotient[quot_thm]: assumes q: "Quotient R Abs Rep" - shows "Quotient (option_rel R) (option_map Abs) (option_map Rep)" - apply(unfold Quotient_def) - apply(auto) - apply(case_tac a, simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q]) - apply(case_tac a, simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q]) - apply(case_tac [!] r) - apply(case_tac [!] s) - apply(simp_all add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q] ) + shows "Quotient (option_rel R) (Option.map Abs) (Option.map Rep)" + unfolding Quotient_def + apply(simp add: split_option_all) + apply(simp add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q]) using q unfolding Quotient_def - apply(blast)+ + apply(blast) done lemma option_equivp[quot_equiv]: @@ -39,16 +41,10 @@ shows "equivp (option_rel R)" apply(rule equivpI) unfolding reflp_def symp_def transp_def - apply(auto) - apply(case_tac [!] x) - apply(simp_all add: equivp_reflp[OF a]) - apply(case_tac [!] y) - apply(simp_all add: equivp_symp[OF a]) - apply(case_tac [!] z) - apply(simp_all) - apply(clarify) - apply(rule equivp_transp[OF a]) - apply(assumption)+ + apply(simp_all add: split_option_all) + apply(blast intro: equivp_reflp[OF a]) + apply(blast intro: equivp_symp[OF a]) + apply(blast intro: equivp_transp[OF a]) done lemma option_None_rsp[quot_respect]: @@ -63,29 +59,22 @@ lemma option_None_prs[quot_preserve]: assumes q: "Quotient R Abs Rep" - shows "option_map Abs None = None" + shows "Option.map Abs None = None" by simp lemma option_Some_prs[quot_respect]: assumes q: "Quotient R Abs Rep" - shows "(Rep ---> option_map Abs) Some = Some" + shows "(Rep ---> Option.map Abs) Some = Some" apply(simp add: expand_fun_eq) apply(simp add: Quotient_abs_rep[OF q]) done -lemma option_map_id[id_simps]: - shows "option_map id = id" - apply (auto simp add: expand_fun_eq) - apply (case_tac x) - apply (auto) - done +lemma option_map_id[id_simps]: + shows "Option.map id = id" + by (simp add: expand_fun_eq split_option_all) -lemma option_rel_eq[id_simps]: +lemma option_rel_eq[id_simps]: shows "option_rel (op =) = (op =)" - apply(auto simp add: expand_fun_eq) - apply(case_tac x) - apply(case_tac [!] xa) - apply(auto) - done + by (simp add: expand_fun_eq split_option_all) end diff -r 0879d144aaa3 -r da5e4b8317c7 Quot/QuotProd.thy --- a/Quot/QuotProd.thy Tue Jan 26 09:54:43 2010 +0100 +++ b/Quot/QuotProd.thy Tue Jan 26 10:53:44 2010 +0100 @@ -1,11 +1,16 @@ +(* Title: QuotProd.thy + Author: Cezary Kaliszyk and Christian Urban +*) theory QuotProd imports QuotMain begin +section {* Quotient infrastructure for product type *} + fun prod_rel where - "prod_rel R1 R2 = (\(a,b) (c,d). R1 a c \ R2 b d)" + "prod_rel R1 R2 = (\(a, b) (c, d). R1 a c \ R2 b d)" declare [[map * = (prod_fun, prod_rel)]] @@ -14,12 +19,12 @@ assumes a: "equivp R1" assumes b: "equivp R2" shows "equivp (prod_rel R1 R2)" - unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def - apply (auto simp add: equivp_reflp[OF a] equivp_reflp[OF b]) - apply (simp only: equivp_symp[OF a]) - apply (simp only: equivp_symp[OF b]) - using equivp_transp[OF a] apply blast - using equivp_transp[OF b] apply blast + apply(rule equivpI) + unfolding reflp_def symp_def transp_def + apply(simp_all add: split_paired_all) + apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b]) + apply(blast intro: equivp_symp[OF a] equivp_symp[OF b]) + apply(blast intro: equivp_transp[OF a] equivp_transp[OF b]) done lemma prod_quotient[quot_thm]: @@ -27,23 +32,26 @@ assumes q2: "Quotient R2 Abs2 Rep2" shows "Quotient (prod_rel R1 R2) (prod_fun Abs1 Abs2) (prod_fun Rep1 Rep2)" unfolding Quotient_def + apply(simp add: split_paired_all) + apply(simp add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1]) + apply(simp add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2]) using q1 q2 - apply (simp add: Quotient_abs_rep Quotient_rel_rep) - using Quotient_rel[OF q1] Quotient_rel[OF q2] - by blast + unfolding Quotient_def + apply(blast) + done -lemma pair_rsp[quot_respect]: +lemma Pair_rsp[quot_respect]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair" - by simp +by simp -lemma pair_prs[quot_preserve]: +lemma Pair_prs[quot_preserve]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(Rep1 ---> Rep2 ---> (prod_fun Abs1 Abs2)) Pair = Pair" - apply (simp add: expand_fun_eq) - apply (simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]) + apply(simp add: expand_fun_eq) + apply(simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]) done lemma fst_rsp[quot_respect]: @@ -56,8 +64,8 @@ assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(prod_fun Rep1 Rep2 ---> Abs1) fst = fst" - apply (simp add: expand_fun_eq) - apply (simp add: Quotient_abs_rep[OF q1]) + apply(simp add: expand_fun_eq) + apply(simp add: Quotient_abs_rep[OF q1]) done lemma snd_rsp[quot_respect]: @@ -65,33 +73,32 @@ assumes "Quotient R2 Abs2 Rep2" shows "(prod_rel R1 R2 ===> R2) snd snd" by simp - + lemma snd_prs[quot_preserve]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "(prod_fun Rep1 Rep2 ---> Abs2) snd = snd" - apply (simp add: expand_fun_eq) - apply (simp add: Quotient_abs_rep[OF q2]) + apply(simp add: expand_fun_eq) + apply(simp add: Quotient_abs_rep[OF q2]) done lemma split_rsp[quot_respect]: - "((R1 ===> R2 ===> op =) ===> (prod_rel R1 R2) ===> op =) split split" + shows "((R1 ===> R2 ===> (op =)) ===> (prod_rel R1 R2) ===> (op =)) split split" by auto - + lemma split_prs[quot_preserve]: assumes q1: "Quotient R1 Abs1 Rep1" and q2: "Quotient R2 Abs2 Rep2" shows "(((Abs1 ---> Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> id) split) = split" by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]) -lemma prod_fun_id[id_simps]: +lemma prod_fun_id[id_simps]: shows "prod_fun id id = id" by (simp add: prod_fun_def) -lemma prod_rel_eq[id_simps]: - shows "(prod_rel (op =) (op =)) = (op =)" - apply (rule ext)+ - apply auto - done +lemma prod_rel_eq[id_simps]: + shows "prod_rel (op =) (op =) = (op =)" + by (simp add: expand_fun_eq) + end diff -r 0879d144aaa3 -r da5e4b8317c7 Quot/QuotSum.thy --- a/Quot/QuotSum.thy Tue Jan 26 09:54:43 2010 +0100 +++ b/Quot/QuotSum.thy Tue Jan 26 10:53:44 2010 +0100 @@ -1,3 +1,6 @@ +(* Title: QuotSum.thy + Author: Cezary Kaliszyk and Christian Urban +*) theory QuotSum imports QuotMain begin @@ -19,79 +22,73 @@ declare [[map "+" = (sum_map, sum_rel)]] +text {* should probably be in Sum_Type.thy *} +lemma split_sum_all: + shows "(\x. P x) \ (\x. P (Inl x)) \ (\x. P (Inr x))" +apply(auto) +apply(case_tac x) +apply(simp_all) +done + lemma sum_equivp[quot_equiv]: assumes a: "equivp R1" assumes b: "equivp R2" shows "equivp (sum_rel R1 R2)" apply(rule equivpI) unfolding reflp_def symp_def transp_def - apply(auto) - apply(case_tac [!] x) - apply(simp_all add: equivp_reflp[OF a] equivp_reflp[OF b]) - apply(case_tac [!] y) - apply(simp_all add: equivp_symp[OF a] equivp_symp[OF b]) - apply(case_tac [!] z) - apply(simp_all) - apply(rule equivp_transp[OF a]) - apply(assumption)+ - apply(rule equivp_transp[OF b]) - apply(assumption)+ + apply(simp_all add: split_sum_all) + apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b]) + apply(blast intro: equivp_symp[OF a] equivp_symp[OF b]) + apply(blast intro: equivp_transp[OF a] equivp_transp[OF b]) done -(* -lemma sum_fun_fun: - assumes q1: "Quotient R1 Abs1 Rep1" - assumes q2: "Quotient R2 Abs2 Rep2" - shows "sum_rel R1 R2 r s = - (sum_rel R1 R2 r r \ sum_rel R1 R2 s s \ sum_map Abs1 Abs2 r = sum_map Abs1 Abs2 s)" - using q1 q2 - apply(case_tac r) - apply(case_tac s) - apply(simp_all) - prefer 2 - apply(case_tac s) - apply(auto) - unfolding Quotient_def - apply metis+ - done -*) - lemma sum_quotient[quot_thm]: assumes q1: "Quotient R1 Abs1 Rep1" assumes q2: "Quotient R2 Abs2 Rep2" shows "Quotient (sum_rel R1 R2) (sum_map Abs1 Abs2) (sum_map Rep1 Rep2)" unfolding Quotient_def - apply(auto) - apply(case_tac a) - apply(simp_all add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1] - Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2]) - apply(case_tac a) - apply(simp_all add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1] - Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2]) - apply(case_tac [!] r) - apply(case_tac [!] s) - apply(simp_all) + apply(simp add: split_sum_all) + apply(simp_all add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1]) + apply(simp_all add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2]) using q1 q2 unfolding Quotient_def apply(blast)+ done -lemma sum_map_id[id_simps]: - shows "sum_map id id = id" - apply (rule ext) - apply (case_tac x) - apply (auto) +lemma sum_Inl_rsp[quot_respect]: + assumes q1: "Quotient R1 Abs1 Rep1" + assumes q2: "Quotient R2 Abs2 Rep2" + shows "(R1 ===> sum_rel R1 R2) Inl Inl" + by simp + +lemma sum_Inr_rsp[quot_respect]: + assumes q1: "Quotient R1 Abs1 Rep1" + assumes q2: "Quotient R2 Abs2 Rep2" + shows "(R2 ===> sum_rel R1 R2) Inr Inr" + by simp + +lemma sum_Inl_prs[quot_respect]: + assumes q1: "Quotient R1 Abs1 Rep1" + assumes q2: "Quotient R2 Abs2 Rep2" + shows "(Rep1 ---> sum_map Abs1 Abs2) Inl = Inl" + apply(simp add: expand_fun_eq) + apply(simp add: Quotient_abs_rep[OF q1]) done -lemma sum_rel_eq[id_simps]: - "(sum_rel op = op =) = op =" - apply (rule ext)+ - apply (case_tac x) - apply auto - apply (case_tac xa) - apply auto - apply (case_tac xa) - apply auto +lemma sum_Inr_prs[quot_respect]: + assumes q1: "Quotient R1 Abs1 Rep1" + assumes q2: "Quotient R2 Abs2 Rep2" + shows "(Rep2 ---> sum_map Abs1 Abs2) Inr = Inr" + apply(simp add: expand_fun_eq) + apply(simp add: Quotient_abs_rep[OF q2]) done +lemma sum_map_id[id_simps]: + shows "sum_map id id = id" + by (simp add: expand_fun_eq split_sum_all) + +lemma sum_rel_eq[id_simps]: + shows "sum_rel (op =) (op =) = (op =)" + by (simp add: expand_fun_eq split_sum_all) + end