diff -r 762f0eae88fd -r 0b15b83ded4a Quot/QuotSum.thy --- a/Quot/QuotSum.thy Tue Jan 26 00:47:40 2010 +0100 +++ b/Quot/QuotSum.thy Tue Jan 26 01:00:35 2010 +0100 @@ -1,3 +1,6 @@ +(* Title: QuotSum.thy + Author: Cezary Kaliszyk and Christian Urban +*) theory QuotSum imports QuotMain begin @@ -19,81 +22,77 @@ 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_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_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" apply (rule eq_reflection) - apply (rule ext) - apply (case_tac x) - apply (auto) + apply (simp add: expand_fun_eq split_sum_all) done lemma sum_rel_eq[id_simps]: - "sum_rel op = op = \ op =" + shows "sum_rel (op =) (op =) \ (op =)" apply (rule eq_reflection) - apply (rule ext)+ - apply (case_tac x) - apply auto - apply (case_tac xa) - apply auto - apply (case_tac xa) - apply auto + apply(simp add: expand_fun_eq split_sum_all) done end