| author | wenzelm | 
| Thu, 05 Sep 2013 23:14:28 +0200 | |
| changeset 53425 | f5b1f555b73b | 
| parent 53026 | e1a548c11845 | 
| child 55564 | e81ee43ab290 | 
| permissions | -rw-r--r-- | 
| 47455 | 1 | (* Title: HOL/Library/Quotient_Sum.thy | 
| 53012 
cb82606b8215
move Lifting/Transfer relevant parts of Library/Quotient_* to Main
 kuncar parents: 
53010diff
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 | |
| 5 | header {* Quotient infrastructure for the sum type *}
 | |
| 6 | ||
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 7 | theory Quotient_Sum | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 8 | imports Main Quotient_Syntax | 
| 
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 | |
| 53012 
cb82606b8215
move Lifting/Transfer relevant parts of Library/Quotient_* to Main
 kuncar parents: 
53010diff
changeset | 11 | subsection {* Rules for the Quotient package *}
 | 
| 51956 
a4d81cdebf8b
better support for domains in Lifting/Transfer = replace Domainp T by the actual invariant in a transferred goal
 kuncar parents: 
51377diff
changeset | 12 | |
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 13 | lemma sum_rel_map1: | 
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 14 | "sum_rel R1 R2 (sum_map f1 f2 x) y \<longleftrightarrow> sum_rel (\<lambda>x. R1 (f1 x)) (\<lambda>x. R2 (f2 x)) x y" | 
| 53026 
e1a548c11845
got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
 traytel parents: 
53012diff
changeset | 15 | by (simp add: sum_rel_def split: sum.split) | 
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 16 | |
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 17 | lemma sum_rel_map2: | 
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 18 | "sum_rel R1 R2 x (sum_map f1 f2 y) \<longleftrightarrow> sum_rel (\<lambda>x y. R1 x (f1 y)) (\<lambda>x y. R2 x (f2 y)) x y" | 
| 53026 
e1a548c11845
got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
 traytel parents: 
53012diff
changeset | 19 | by (simp add: sum_rel_def split: sum.split) | 
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 20 | |
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 21 | lemma sum_map_id [id_simps]: | 
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 22 | "sum_map id id = id" | 
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 23 | by (simp add: id_def sum_map.identity fun_eq_iff) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 24 | |
| 53012 
cb82606b8215
move Lifting/Transfer relevant parts of Library/Quotient_* to Main
 kuncar parents: 
53010diff
changeset | 25 | lemma sum_rel_eq [id_simps]: | 
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 26 | "sum_rel (op =) (op =) = (op =)" | 
| 53026 
e1a548c11845
got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
 traytel parents: 
53012diff
changeset | 27 | by (simp add: sum_rel_def fun_eq_iff split: sum.split) | 
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 28 | |
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 29 | lemma sum_symp: | 
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 30 | "symp R1 \<Longrightarrow> symp R2 \<Longrightarrow> symp (sum_rel R1 R2)" | 
| 53026 
e1a548c11845
got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
 traytel parents: 
53012diff
changeset | 31 | unfolding symp_def split_sum_all sum_rel_simps by fast | 
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 32 | |
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 33 | lemma sum_transp: | 
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 34 | "transp R1 \<Longrightarrow> transp R2 \<Longrightarrow> transp (sum_rel R1 R2)" | 
| 53026 
e1a548c11845
got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
 traytel parents: 
53012diff
changeset | 35 | unfolding transp_def split_sum_all sum_rel_simps by fast | 
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 36 | |
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 37 | lemma sum_equivp [quot_equiv]: | 
| 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 38 | "equivp R1 \<Longrightarrow> equivp R2 \<Longrightarrow> equivp (sum_rel R1 R2)" | 
| 51994 | 39 | by (blast intro: equivpI reflp_sum_rel sum_symp sum_transp elim: equivpE) | 
| 47624 
16d433895d2e
add new transfer rules and setup for lifting package
 huffman parents: 
47455diff
changeset | 40 | |
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 41 | lemma sum_quotient [quot_thm]: | 
| 47308 | 42 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 43 | assumes q2: "Quotient3 R2 Abs2 Rep2" | |
| 44 | shows "Quotient3 (sum_rel R1 R2) (sum_map Abs1 Abs2) (sum_map Rep1 Rep2)" | |
| 45 | apply (rule Quotient3I) | |
| 41372 | 46 | apply (simp_all add: sum_map.compositionality comp_def sum_map.identity sum_rel_eq sum_rel_map1 sum_rel_map2 | 
| 47308 | 47 | Quotient3_abs_rep [OF q1] Quotient3_rel_rep [OF q1] Quotient3_abs_rep [OF q2] Quotient3_rel_rep [OF q2]) | 
| 48 | using Quotient3_rel [OF q1] Quotient3_rel [OF q2] | |
| 53026 
e1a548c11845
got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
 traytel parents: 
53012diff
changeset | 49 | apply (simp add: sum_rel_def comp_def split: sum.split) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 50 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 51 | |
| 47308 | 52 | declare [[mapQ3 sum = (sum_rel, sum_quotient)]] | 
| 47094 | 53 | |
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 54 | lemma sum_Inl_rsp [quot_respect]: | 
| 47308 | 55 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 56 | assumes q2: "Quotient3 R2 Abs2 Rep2" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 57 | shows "(R1 ===> sum_rel R1 R2) Inl Inl" | 
| 40465 
2989f9f3aa10
more appropriate specification packages; fun_rel_def is no simp rule by default
 haftmann parents: 
39302diff
changeset | 58 | by auto | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 59 | |
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 60 | lemma sum_Inr_rsp [quot_respect]: | 
| 47308 | 61 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 62 | assumes q2: "Quotient3 R2 Abs2 Rep2" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 63 | shows "(R2 ===> sum_rel R1 R2) Inr Inr" | 
| 40465 
2989f9f3aa10
more appropriate specification packages; fun_rel_def is no simp rule by default
 haftmann parents: 
39302diff
changeset | 64 | by auto | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 65 | |
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 66 | lemma sum_Inl_prs [quot_preserve]: | 
| 47308 | 67 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 68 | assumes q2: "Quotient3 R2 Abs2 Rep2" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 69 | shows "(Rep1 ---> sum_map Abs1 Abs2) Inl = Inl" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 70 | apply(simp add: fun_eq_iff) | 
| 47308 | 71 | apply(simp add: Quotient3_abs_rep[OF q1]) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 72 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 73 | |
| 40820 
fd9c98ead9a9
more systematic and compact proofs on type relation operators using natural deduction rules
 haftmann parents: 
40610diff
changeset | 74 | lemma sum_Inr_prs [quot_preserve]: | 
| 47308 | 75 | assumes q1: "Quotient3 R1 Abs1 Rep1" | 
| 76 | assumes q2: "Quotient3 R2 Abs2 Rep2" | |
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 77 | shows "(Rep2 ---> sum_map Abs1 Abs2) Inr = Inr" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 78 | apply(simp add: fun_eq_iff) | 
| 47308 | 79 | apply(simp add: Quotient3_abs_rep[OF q2]) | 
| 35222 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 80 | done | 
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 81 | |
| 
4f1fba00f66d
Initial version of HOL quotient package.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 82 | end |