1 (* Title: Quotient_Sum.thy |
|
2 Author: Cezary Kaliszyk and Christian Urban |
|
3 *) |
|
4 theory Quotient_Sum |
|
5 imports Quotient Quotient_Syntax |
|
6 begin |
|
7 |
|
8 section {* Quotient infrastructure for the sum type. *} |
|
9 |
|
10 fun |
|
11 sum_rel |
|
12 where |
|
13 "sum_rel R1 R2 (Inl a1) (Inl b1) = R1 a1 b1" |
|
14 | "sum_rel R1 R2 (Inl a1) (Inr b2) = False" |
|
15 | "sum_rel R1 R2 (Inr a2) (Inl b1) = False" |
|
16 | "sum_rel R1 R2 (Inr a2) (Inr b2) = R2 a2 b2" |
|
17 |
|
18 fun |
|
19 sum_map |
|
20 where |
|
21 "sum_map f1 f2 (Inl a) = Inl (f1 a)" |
|
22 | "sum_map f1 f2 (Inr a) = Inr (f2 a)" |
|
23 |
|
24 declare [[map "+" = (sum_map, sum_rel)]] |
|
25 |
|
26 |
|
27 text {* should probably be in Sum_Type.thy *} |
|
28 lemma split_sum_all: |
|
29 shows "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (Inl x)) \<and> (\<forall>x. P (Inr x))" |
|
30 apply(auto) |
|
31 apply(case_tac x) |
|
32 apply(simp_all) |
|
33 done |
|
34 |
|
35 lemma sum_equivp[quot_equiv]: |
|
36 assumes a: "equivp R1" |
|
37 assumes b: "equivp R2" |
|
38 shows "equivp (sum_rel R1 R2)" |
|
39 apply(rule equivpI) |
|
40 unfolding reflp_def symp_def transp_def |
|
41 apply(simp_all add: split_sum_all) |
|
42 apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b]) |
|
43 apply(blast intro: equivp_symp[OF a] equivp_symp[OF b]) |
|
44 apply(blast intro: equivp_transp[OF a] equivp_transp[OF b]) |
|
45 done |
|
46 |
|
47 lemma sum_quotient[quot_thm]: |
|
48 assumes q1: "Quotient R1 Abs1 Rep1" |
|
49 assumes q2: "Quotient R2 Abs2 Rep2" |
|
50 shows "Quotient (sum_rel R1 R2) (sum_map Abs1 Abs2) (sum_map Rep1 Rep2)" |
|
51 unfolding Quotient_def |
|
52 apply(simp add: split_sum_all) |
|
53 apply(simp_all add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1]) |
|
54 apply(simp_all add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2]) |
|
55 using q1 q2 |
|
56 unfolding Quotient_def |
|
57 apply(blast)+ |
|
58 done |
|
59 |
|
60 lemma sum_Inl_rsp[quot_respect]: |
|
61 assumes q1: "Quotient R1 Abs1 Rep1" |
|
62 assumes q2: "Quotient R2 Abs2 Rep2" |
|
63 shows "(R1 ===> sum_rel R1 R2) Inl Inl" |
|
64 by simp |
|
65 |
|
66 lemma sum_Inr_rsp[quot_respect]: |
|
67 assumes q1: "Quotient R1 Abs1 Rep1" |
|
68 assumes q2: "Quotient R2 Abs2 Rep2" |
|
69 shows "(R2 ===> sum_rel R1 R2) Inr Inr" |
|
70 by simp |
|
71 |
|
72 lemma sum_Inl_prs[quot_preserve]: |
|
73 assumes q1: "Quotient R1 Abs1 Rep1" |
|
74 assumes q2: "Quotient R2 Abs2 Rep2" |
|
75 shows "(Rep1 ---> sum_map Abs1 Abs2) Inl = Inl" |
|
76 apply(simp add: expand_fun_eq) |
|
77 apply(simp add: Quotient_abs_rep[OF q1]) |
|
78 done |
|
79 |
|
80 lemma sum_Inr_prs[quot_preserve]: |
|
81 assumes q1: "Quotient R1 Abs1 Rep1" |
|
82 assumes q2: "Quotient R2 Abs2 Rep2" |
|
83 shows "(Rep2 ---> sum_map Abs1 Abs2) Inr = Inr" |
|
84 apply(simp add: expand_fun_eq) |
|
85 apply(simp add: Quotient_abs_rep[OF q2]) |
|
86 done |
|
87 |
|
88 lemma sum_map_id[id_simps]: |
|
89 shows "sum_map id id = id" |
|
90 by (simp add: expand_fun_eq split_sum_all) |
|
91 |
|
92 lemma sum_rel_eq[id_simps]: |
|
93 shows "sum_rel (op =) (op =) = (op =)" |
|
94 by (simp add: expand_fun_eq split_sum_all) |
|
95 |
|
96 end |
|