| author | wenzelm | 
| Fri, 28 Apr 2017 13:21:03 +0200 | |
| changeset 65604 | 637aa8e93cd7 | 
| parent 62954 | c5d0fdc260fa | 
| child 67399 | eab6ce8368fa | 
| permissions | -rw-r--r-- | 
| 47455 | 1  | 
(* Title: HOL/Library/Quotient_Product.thy  | 
| 
53012
 
cb82606b8215
move Lifting/Transfer relevant parts of Library/Quotient_* to Main
 
kuncar 
parents: 
51994 
diff
changeset
 | 
2  | 
Author: Cezary Kaliszyk and Christian Urban  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
3  | 
*)  | 
| 35788 | 4  | 
|
| 60500 | 5  | 
section \<open>Quotient infrastructure for the product type\<close>  | 
| 35788 | 6  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
7  | 
theory Quotient_Product  | 
| 62954 | 8  | 
imports Quotient_Syntax  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
9  | 
begin  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
10  | 
|
| 60500 | 11  | 
subsection \<open>Rules for the Quotient package\<close>  | 
| 
51956
 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 
kuncar 
parents: 
51377 
diff
changeset
 | 
12  | 
|
| 55932 | 13  | 
lemma map_prod_id [id_simps]:  | 
14  | 
shows "map_prod id id = id"  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
15  | 
by (simp add: fun_eq_iff)  | 
| 
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
16  | 
|
| 55944 | 17  | 
lemma rel_prod_eq [id_simps]:  | 
18  | 
shows "rel_prod (op =) (op =) = (op =)"  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
19  | 
by (simp add: fun_eq_iff)  | 
| 
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
20  | 
|
| 
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
21  | 
lemma prod_equivp [quot_equiv]:  | 
| 
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
22  | 
assumes "equivp R1"  | 
| 
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
23  | 
assumes "equivp R2"  | 
| 55944 | 24  | 
shows "equivp (rel_prod R1 R2)"  | 
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
25  | 
using assms by (auto intro!: equivpI reflpI sympI transpI elim!: equivpE elim: reflpE sympE transpE)  | 
| 
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
26  | 
|
| 
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
27  | 
lemma prod_quotient [quot_thm]:  | 
| 47308 | 28  | 
assumes "Quotient3 R1 Abs1 Rep1"  | 
29  | 
assumes "Quotient3 R2 Abs2 Rep2"  | 
|
| 55944 | 30  | 
shows "Quotient3 (rel_prod R1 R2) (map_prod Abs1 Abs2) (map_prod Rep1 Rep2)"  | 
| 47308 | 31  | 
apply (rule Quotient3I)  | 
| 55932 | 32  | 
apply (simp add: map_prod.compositionality comp_def map_prod.identity  | 
| 47308 | 33  | 
Quotient3_abs_rep [OF assms(1)] Quotient3_abs_rep [OF assms(2)])  | 
34  | 
apply (simp add: split_paired_all Quotient3_rel_rep [OF assms(1)] Quotient3_rel_rep [OF assms(2)])  | 
|
35  | 
using Quotient3_rel [OF assms(1)] Quotient3_rel [OF assms(2)]  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
36  | 
apply (auto simp add: split_paired_all)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
37  | 
done  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
38  | 
|
| 55944 | 39  | 
declare [[mapQ3 prod = (rel_prod, prod_quotient)]]  | 
| 47094 | 40  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
41  | 
lemma Pair_rsp [quot_respect]:  | 
| 47308 | 42  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
43  | 
assumes q2: "Quotient3 R2 Abs2 Rep2"  | 
|
| 55944 | 44  | 
shows "(R1 ===> R2 ===> rel_prod R1 R2) Pair Pair"  | 
| 
47624
 
16d433895d2e
add new transfer rules and setup for lifting package
 
huffman 
parents: 
47455 
diff
changeset
 | 
45  | 
by (rule Pair_transfer)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
46  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
47  | 
lemma Pair_prs [quot_preserve]:  | 
| 47308 | 48  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
49  | 
assumes q2: "Quotient3 R2 Abs2 Rep2"  | 
|
| 55932 | 50  | 
shows "(Rep1 ---> Rep2 ---> (map_prod Abs1 Abs2)) Pair = Pair"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
51  | 
apply(simp add: fun_eq_iff)  | 
| 47308 | 52  | 
apply(simp add: Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2])  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
53  | 
done  | 
| 
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
54  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
55  | 
lemma fst_rsp [quot_respect]:  | 
| 47308 | 56  | 
assumes "Quotient3 R1 Abs1 Rep1"  | 
57  | 
assumes "Quotient3 R2 Abs2 Rep2"  | 
|
| 55944 | 58  | 
shows "(rel_prod R1 R2 ===> R1) fst fst"  | 
| 
40465
 
2989f9f3aa10
more appropriate specification packages; fun_rel_def is no simp rule by default
 
haftmann 
parents: 
39302 
diff
changeset
 | 
59  | 
by auto  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
60  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
61  | 
lemma fst_prs [quot_preserve]:  | 
| 47308 | 62  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
63  | 
assumes q2: "Quotient3 R2 Abs2 Rep2"  | 
|
| 55932 | 64  | 
shows "(map_prod Rep1 Rep2 ---> Abs1) fst = fst"  | 
| 47308 | 65  | 
by (simp add: fun_eq_iff Quotient3_abs_rep[OF q1])  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
66  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
67  | 
lemma snd_rsp [quot_respect]:  | 
| 47308 | 68  | 
assumes "Quotient3 R1 Abs1 Rep1"  | 
69  | 
assumes "Quotient3 R2 Abs2 Rep2"  | 
|
| 55944 | 70  | 
shows "(rel_prod R1 R2 ===> R2) snd snd"  | 
| 
40465
 
2989f9f3aa10
more appropriate specification packages; fun_rel_def is no simp rule by default
 
haftmann 
parents: 
39302 
diff
changeset
 | 
71  | 
by auto  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
72  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
73  | 
lemma snd_prs [quot_preserve]:  | 
| 47308 | 74  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
75  | 
assumes q2: "Quotient3 R2 Abs2 Rep2"  | 
|
| 55932 | 76  | 
shows "(map_prod Rep1 Rep2 ---> Abs2) snd = snd"  | 
| 47308 | 77  | 
by (simp add: fun_eq_iff Quotient3_abs_rep[OF q2])  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
78  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
60500 
diff
changeset
 | 
79  | 
lemma case_prod_rsp [quot_respect]:  | 
| 
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
60500 
diff
changeset
 | 
80  | 
shows "((R1 ===> R2 ===> (op =)) ===> (rel_prod R1 R2) ===> (op =)) case_prod case_prod"  | 
| 
55414
 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 
blanchet 
parents: 
53012 
diff
changeset
 | 
81  | 
by (rule case_prod_transfer)  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
82  | 
|
| 
40820
 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 
haftmann 
parents: 
40607 
diff
changeset
 | 
83  | 
lemma split_prs [quot_preserve]:  | 
| 47308 | 84  | 
assumes q1: "Quotient3 R1 Abs1 Rep1"  | 
85  | 
and q2: "Quotient3 R2 Abs2 Rep2"  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
60500 
diff
changeset
 | 
86  | 
shows "(((Abs1 ---> Abs2 ---> id) ---> map_prod Rep1 Rep2 ---> id) case_prod) = case_prod"  | 
| 47308 | 87  | 
by (simp add: fun_eq_iff Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2])  | 
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
88  | 
|
| 
36695
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
89  | 
lemma [quot_respect]:  | 
| 
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
90  | 
shows "((R2 ===> R2 ===> op =) ===> (R1 ===> R1 ===> op =) ===>  | 
| 55944 | 91  | 
rel_prod R2 R1 ===> rel_prod R2 R1 ===> op =) rel_prod rel_prod"  | 
| 58916 | 92  | 
by (rule prod.rel_transfer)  | 
| 
36695
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
93  | 
|
| 
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
94  | 
lemma [quot_preserve]:  | 
| 47308 | 95  | 
assumes q1: "Quotient3 R1 abs1 rep1"  | 
96  | 
and q2: "Quotient3 R2 abs2 rep2"  | 
|
| 
36695
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
97  | 
shows "((abs1 ---> abs1 ---> id) ---> (abs2 ---> abs2 ---> id) --->  | 
| 55944 | 98  | 
map_prod rep1 rep2 ---> map_prod rep1 rep2 ---> id) rel_prod = rel_prod"  | 
| 47308 | 99  | 
by (simp add: fun_eq_iff Quotient3_abs_rep[OF q1] Quotient3_abs_rep[OF q2])  | 
| 
36695
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
100  | 
|
| 
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
101  | 
lemma [quot_preserve]:  | 
| 55944 | 102  | 
shows"(rel_prod ((rep1 ---> rep1 ---> id) R1) ((rep2 ---> rep2 ---> id) R2)  | 
| 
36695
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
103  | 
(l1, l2) (r1, r2)) = (R1 (rep1 l1) (rep1 r1) \<and> R2 (rep2 l2) (rep2 r2))"  | 
| 
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
104  | 
by simp  | 
| 
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
105  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
60500 
diff
changeset
 | 
106  | 
declare prod.inject[quot_preserve]  | 
| 
36695
 
b434506fb0d4
respectfullness and preservation of prod_rel
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents: 
35788 
diff
changeset
 | 
107  | 
|
| 
35222
 
4f1fba00f66d
Initial version of HOL quotient package.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
108  | 
end  |