1128
|
1 |
(* Title: Quotient_Product.thy
|
935
|
2 |
Author: Cezary Kaliszyk and Christian Urban
|
|
3 |
*)
|
1128
|
4 |
theory Quotient_Product
|
1129
|
5 |
imports Quotient Quotient_Syntax
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
begin
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
|
937
|
8 |
section {* Quotient infrastructure for the product type. *}
|
935
|
9 |
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
10 |
fun
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
11 |
prod_rel
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
where
|
935
|
13 |
"prod_rel R1 R2 = (\<lambda>(a, b) (c, d). R1 a c \<and> R2 b d)"
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
14 |
|
922
|
15 |
declare [[map * = (prod_fun, prod_rel)]]
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
16 |
|
922
|
17 |
|
|
18 |
lemma prod_equivp[quot_equiv]:
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
19 |
assumes a: "equivp R1"
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
assumes b: "equivp R2"
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
21 |
shows "equivp (prod_rel R1 R2)"
|
935
|
22 |
apply(rule equivpI)
|
|
23 |
unfolding reflp_def symp_def transp_def
|
|
24 |
apply(simp_all add: split_paired_all)
|
|
25 |
apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b])
|
|
26 |
apply(blast intro: equivp_symp[OF a] equivp_symp[OF b])
|
1128
|
27 |
apply(blast intro: equivp_transp[OF a] equivp_transp[OF b])
|
928
|
28 |
done
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
29 |
|
922
|
30 |
lemma prod_quotient[quot_thm]:
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
31 |
assumes q1: "Quotient R1 Abs1 Rep1"
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
32 |
assumes q2: "Quotient R2 Abs2 Rep2"
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
33 |
shows "Quotient (prod_rel R1 R2) (prod_fun Abs1 Abs2) (prod_fun Rep1 Rep2)"
|
928
|
34 |
unfolding Quotient_def
|
935
|
35 |
apply(simp add: split_paired_all)
|
|
36 |
apply(simp add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1])
|
|
37 |
apply(simp add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2])
|
928
|
38 |
using q1 q2
|
935
|
39 |
unfolding Quotient_def
|
|
40 |
apply(blast)
|
|
41 |
done
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
42 |
|
935
|
43 |
lemma Pair_rsp[quot_respect]:
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
44 |
assumes q1: "Quotient R1 Abs1 Rep1"
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
45 |
assumes q2: "Quotient R2 Abs2 Rep2"
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
46 |
shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair"
|
937
|
47 |
by simp
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
48 |
|
935
|
49 |
lemma Pair_prs[quot_preserve]:
|
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
50 |
assumes q1: "Quotient R1 Abs1 Rep1"
|
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
51 |
assumes q2: "Quotient R2 Abs2 Rep2"
|
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
52 |
shows "(Rep1 ---> Rep2 ---> (prod_fun Abs1 Abs2)) Pair = Pair"
|
935
|
53 |
apply(simp add: expand_fun_eq)
|
|
54 |
apply(simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
|
928
|
55 |
done
|
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
56 |
|
922
|
57 |
lemma fst_rsp[quot_respect]:
|
|
58 |
assumes "Quotient R1 Abs1 Rep1"
|
|
59 |
assumes "Quotient R2 Abs2 Rep2"
|
|
60 |
shows "(prod_rel R1 R2 ===> R1) fst fst"
|
|
61 |
by simp
|
648
830b58c2fa94
decoupled QuotProd from QuotMain and also started new cleaning strategy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
62 |
|
922
|
63 |
lemma fst_prs[quot_preserve]:
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
64 |
assumes q1: "Quotient R1 Abs1 Rep1"
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
65 |
assumes q2: "Quotient R2 Abs2 Rep2"
|
922
|
66 |
shows "(prod_fun Rep1 Rep2 ---> Abs1) fst = fst"
|
935
|
67 |
apply(simp add: expand_fun_eq)
|
|
68 |
apply(simp add: Quotient_abs_rep[OF q1])
|
928
|
69 |
done
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
70 |
|
922
|
71 |
lemma snd_rsp[quot_respect]:
|
|
72 |
assumes "Quotient R1 Abs1 Rep1"
|
|
73 |
assumes "Quotient R2 Abs2 Rep2"
|
|
74 |
shows "(prod_rel R1 R2 ===> R2) snd snd"
|
|
75 |
by simp
|
1128
|
76 |
|
922
|
77 |
lemma snd_prs[quot_preserve]:
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
78 |
assumes q1: "Quotient R1 Abs1 Rep1"
|
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
79 |
assumes q2: "Quotient R2 Abs2 Rep2"
|
922
|
80 |
shows "(prod_fun Rep1 Rep2 ---> Abs2) snd = snd"
|
935
|
81 |
apply(simp add: expand_fun_eq)
|
|
82 |
apply(simp add: Quotient_abs_rep[OF q2])
|
928
|
83 |
done
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
84 |
|
931
|
85 |
lemma split_rsp[quot_respect]:
|
936
|
86 |
shows "((R1 ===> R2 ===> (op =)) ===> (prod_rel R1 R2) ===> (op =)) split split"
|
931
|
87 |
by auto
|
1128
|
88 |
|
931
|
89 |
lemma split_prs[quot_preserve]:
|
|
90 |
assumes q1: "Quotient R1 Abs1 Rep1"
|
|
91 |
and q2: "Quotient R2 Abs2 Rep2"
|
|
92 |
shows "(((Abs1 ---> Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> id) split) = split"
|
|
93 |
by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
|
|
94 |
|
1128
|
95 |
lemma prod_fun_id[id_simps]:
|
928
|
96 |
shows "prod_fun id id = id"
|
|
97 |
by (simp add: prod_fun_def)
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
98 |
|
1128
|
99 |
lemma prod_rel_eq[id_simps]:
|
936
|
100 |
shows "prod_rel (op =) (op =) = (op =)"
|
|
101 |
by (simp add: expand_fun_eq)
|
1128
|
102 |
|
829
42b90994ac77
map and rel simps for all quotients; needed when changing the relations to aggregate ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
|
103 |
|
593
18eac4596ef1
QuotProd with product_quotient and a 3 respects and preserves lemmas.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
104 |
end
|