wenzelm@47455: (* Title: HOL/Library/Quotient_Sum.thy kuncar@53012: Author: Cezary Kaliszyk and Christian Urban kaliszyk@35222: *) wenzelm@35788: wenzelm@35788: header {* Quotient infrastructure for the sum type *} wenzelm@35788: kaliszyk@35222: theory Quotient_Sum kaliszyk@35222: imports Main Quotient_Syntax kaliszyk@35222: begin kaliszyk@35222: kuncar@53012: subsection {* Rules for the Quotient package *} kuncar@51956: blanchet@55943: lemma rel_sum_map1: blanchet@55943: "rel_sum R1 R2 (map_sum f1 f2 x) y \ rel_sum (\x. R1 (f1 x)) (\x. R2 (f2 x)) x y" blanchet@55943: by (simp add: rel_sum_def split: sum.split) haftmann@40820: blanchet@55943: lemma rel_sum_map2: blanchet@55943: "rel_sum R1 R2 x (map_sum f1 f2 y) \ rel_sum (\x y. R1 x (f1 y)) (\x y. R2 x (f2 y)) x y" blanchet@55943: by (simp add: rel_sum_def split: sum.split) haftmann@40820: blanchet@55931: lemma map_sum_id [id_simps]: blanchet@55931: "map_sum id id = id" blanchet@55931: by (simp add: id_def map_sum.identity fun_eq_iff) kaliszyk@35222: blanchet@55943: lemma rel_sum_eq [id_simps]: blanchet@55943: "rel_sum (op =) (op =) = (op =)" blanchet@55943: by (simp add: rel_sum_def fun_eq_iff split: sum.split) haftmann@40820: blanchet@55943: lemma reflp_rel_sum: blanchet@55943: "reflp R1 \ reflp R2 \ reflp (rel_sum R1 R2)" blanchet@55943: unfolding reflp_def split_sum_all rel_sum_simps by fast kuncar@55564: haftmann@40820: lemma sum_symp: blanchet@55943: "symp R1 \ symp R2 \ symp (rel_sum R1 R2)" blanchet@55943: unfolding symp_def split_sum_all rel_sum_simps by fast haftmann@40820: haftmann@40820: lemma sum_transp: blanchet@55943: "transp R1 \ transp R2 \ transp (rel_sum R1 R2)" blanchet@55943: unfolding transp_def split_sum_all rel_sum_simps by fast haftmann@40820: haftmann@40820: lemma sum_equivp [quot_equiv]: blanchet@55943: "equivp R1 \ equivp R2 \ equivp (rel_sum R1 R2)" blanchet@55943: by (blast intro: equivpI reflp_rel_sum sum_symp sum_transp elim: equivpE) huffman@47624: haftmann@40820: lemma sum_quotient [quot_thm]: kuncar@47308: assumes q1: "Quotient3 R1 Abs1 Rep1" kuncar@47308: assumes q2: "Quotient3 R2 Abs2 Rep2" blanchet@55943: shows "Quotient3 (rel_sum R1 R2) (map_sum Abs1 Abs2) (map_sum Rep1 Rep2)" kuncar@47308: apply (rule Quotient3I) blanchet@55943: apply (simp_all add: map_sum.compositionality comp_def map_sum.identity rel_sum_eq rel_sum_map1 rel_sum_map2 kuncar@47308: Quotient3_abs_rep [OF q1] Quotient3_rel_rep [OF q1] Quotient3_abs_rep [OF q2] Quotient3_rel_rep [OF q2]) kuncar@47308: using Quotient3_rel [OF q1] Quotient3_rel [OF q2] blanchet@55943: apply (simp add: rel_sum_def comp_def split: sum.split) kaliszyk@35222: done kaliszyk@35222: blanchet@55943: declare [[mapQ3 sum = (rel_sum, sum_quotient)]] kuncar@47094: haftmann@40820: lemma sum_Inl_rsp [quot_respect]: kuncar@47308: assumes q1: "Quotient3 R1 Abs1 Rep1" kuncar@47308: assumes q2: "Quotient3 R2 Abs2 Rep2" blanchet@55943: shows "(R1 ===> rel_sum R1 R2) Inl Inl" haftmann@40465: by auto kaliszyk@35222: haftmann@40820: lemma sum_Inr_rsp [quot_respect]: kuncar@47308: assumes q1: "Quotient3 R1 Abs1 Rep1" kuncar@47308: assumes q2: "Quotient3 R2 Abs2 Rep2" blanchet@55943: shows "(R2 ===> rel_sum R1 R2) Inr Inr" haftmann@40465: by auto kaliszyk@35222: haftmann@40820: lemma sum_Inl_prs [quot_preserve]: kuncar@47308: assumes q1: "Quotient3 R1 Abs1 Rep1" kuncar@47308: assumes q2: "Quotient3 R2 Abs2 Rep2" blanchet@55931: shows "(Rep1 ---> map_sum Abs1 Abs2) Inl = Inl" nipkow@39302: apply(simp add: fun_eq_iff) kuncar@47308: apply(simp add: Quotient3_abs_rep[OF q1]) kaliszyk@35222: done kaliszyk@35222: haftmann@40820: lemma sum_Inr_prs [quot_preserve]: kuncar@47308: assumes q1: "Quotient3 R1 Abs1 Rep1" kuncar@47308: assumes q2: "Quotient3 R2 Abs2 Rep2" blanchet@55931: shows "(Rep2 ---> map_sum Abs1 Abs2) Inr = Inr" nipkow@39302: apply(simp add: fun_eq_iff) kuncar@47308: apply(simp add: Quotient3_abs_rep[OF q2]) kaliszyk@35222: done kaliszyk@35222: kaliszyk@35222: end