| author | Christian Urban <urbanc@in.tum.de> | 
| Mon, 05 Mar 2012 16:27:28 +0000 | |
| changeset 3132 | 87eca760dcba | 
| parent 1260 | 9df6144e281b | 
| permissions | -rw-r--r-- | 
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 1 | (* Title: Quotient_List.thy | 
| 937 | 2 | Author: Cezary Kaliszyk and Christian Urban | 
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 3 | *) | 
| 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 4 | theory Quotient_List | 
| 1129 
9a86f0ef6503
Notation available locally
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1128diff
changeset | 5 | imports Quotient Quotient_Syntax List | 
| 0 | 6 | begin | 
| 7 | ||
| 937 | 8 | section {* Quotient infrastructure for the list type. *}
 | 
| 9 | ||
| 0 | 10 | fun | 
| 537 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 11 | list_rel | 
| 0 | 12 | where | 
| 537 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 13 | "list_rel R [] [] = True" | 
| 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 14 | | "list_rel R (x#xs) [] = False" | 
| 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 15 | | "list_rel R [] (x#xs) = False" | 
| 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 16 | | "list_rel R (x#xs) (y#ys) = (R x y \<and> list_rel R xs ys)" | 
| 0 | 17 | |
| 600 
5d932e7a856c
List moved after QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
597diff
changeset | 18 | declare [[map list = (map, list_rel)]] | 
| 
5d932e7a856c
List moved after QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
597diff
changeset | 19 | |
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 20 | lemma split_list_all: | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 21 | shows "(\<forall>x. P x) \<longleftrightarrow> P [] \<and> (\<forall>x xs. P (x#xs))" | 
| 936 | 22 | apply(auto) | 
| 23 | apply(case_tac x) | |
| 24 | apply(simp_all) | |
| 25 | done | |
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 26 | |
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 27 | lemma map_id[id_simps]: | 
| 936 | 28 | shows "map id = id" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 29 | apply(simp add: expand_fun_eq) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 30 | apply(rule allI) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 31 | apply(induct_tac x) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 32 | apply(simp_all) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 33 | done | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 34 | |
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 35 | |
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 36 | lemma list_rel_reflp: | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 37 | shows "equivp R \<Longrightarrow> list_rel R xs xs" | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 38 | apply(induct xs) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 39 | apply(simp_all add: equivp_reflp) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 40 | done | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 41 | |
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 42 | lemma list_rel_symp: | 
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 43 | assumes a: "equivp R" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 44 | shows "list_rel R xs ys \<Longrightarrow> list_rel R ys xs" | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 45 | apply(induct xs ys rule: list_induct2') | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 46 | apply(simp_all) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 47 | apply(rule equivp_symp[OF a]) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 48 | apply(simp) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 49 | done | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 50 | |
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 51 | lemma list_rel_transp: | 
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 52 | assumes a: "equivp R" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 53 | shows "list_rel R xs1 xs2 \<Longrightarrow> list_rel R xs2 xs3 \<Longrightarrow> list_rel R xs1 xs3" | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 54 | apply(induct xs1 xs2 arbitrary: xs3 rule: list_induct2') | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 55 | apply(simp_all) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 56 | apply(case_tac xs3) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 57 | apply(simp_all) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 58 | apply(rule equivp_transp[OF a]) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 59 | apply(auto) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 60 | done | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 61 | |
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
623diff
changeset | 62 | lemma list_equivp[quot_equiv]: | 
| 529 | 63 | assumes a: "equivp R" | 
| 537 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 64 | shows "equivp (list_rel R)" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 65 | apply(rule equivpI) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 66 | unfolding reflp_def symp_def transp_def | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 67 | apply(subst split_list_all) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 68 | apply(simp add: equivp_reflp[OF a] list_rel_reflp[OF a]) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 69 | apply(blast intro: list_rel_symp[OF a]) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 70 | apply(blast intro: list_rel_transp[OF a]) | 
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 71 | done | 
| 0 | 72 | |
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 73 | lemma list_rel_rel: | 
| 529 | 74 | assumes q: "Quotient R Abs Rep" | 
| 537 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 75 | shows "list_rel R r s = (list_rel R r r \<and> list_rel R s s \<and> (map Abs r = map Abs s))" | 
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 76 | apply(induct r s rule: list_induct2') | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 77 | apply(simp_all) | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 78 | using Quotient_rel[OF q] | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 79 | apply(metis) | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 80 | done | 
| 0 | 81 | |
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
623diff
changeset | 82 | lemma list_quotient[quot_thm]: | 
| 529 | 83 | assumes q: "Quotient R Abs Rep" | 
| 537 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 84 | shows "Quotient (list_rel R) (map Abs) (map Rep)" | 
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 85 | unfolding Quotient_def | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 86 | apply(subst split_list_all) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 87 | apply(simp add: Quotient_abs_rep[OF q] abs_o_rep[OF q] map_id) | 
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 88 | apply(rule conjI) | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 89 | apply(rule allI) | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 90 | apply(induct_tac a) | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 91 | apply(simp) | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 92 | apply(simp) | 
| 541 
94deffed619d
more name cleaning and removing
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
540diff
changeset | 93 | apply(simp add: Quotient_rep_reflp[OF q]) | 
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 94 | apply(rule allI)+ | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 95 | apply(rule list_rel_rel[OF q]) | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 96 | done | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 97 | |
| 600 
5d932e7a856c
List moved after QuotMain
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
597diff
changeset | 98 | |
| 644 
97a397ba5743
started to reformulate preserve lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 99 | lemma cons_prs_aux: | 
| 540 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 100 | assumes q: "Quotient R Abs Rep" | 
| 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 101 | shows "(map Abs) ((Rep h) # (map Rep t)) = h # t" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 102 | by (induct t) (simp_all add: Quotient_abs_rep[OF q]) | 
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 103 | |
| 644 
97a397ba5743
started to reformulate preserve lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 104 | lemma cons_prs[quot_preserve]: | 
| 
97a397ba5743
started to reformulate preserve lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 105 | assumes q: "Quotient R Abs Rep" | 
| 
97a397ba5743
started to reformulate preserve lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 106 | shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 107 | by (simp only: expand_fun_eq fun_map_def cons_prs_aux[OF q]) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 108 | (simp) | 
| 644 
97a397ba5743
started to reformulate preserve lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 109 | |
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
623diff
changeset | 110 | lemma cons_rsp[quot_respect]: | 
| 529 | 111 | assumes q: "Quotient R Abs Rep" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 112 | shows "(R ===> list_rel R ===> list_rel R) (op #) (op #)" | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 113 | by (auto) | 
| 0 | 114 | |
| 644 
97a397ba5743
started to reformulate preserve lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
636diff
changeset | 115 | lemma nil_prs[quot_preserve]: | 
| 529 | 116 | assumes q: "Quotient R Abs Rep" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 117 | shows "map Abs [] = []" | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 118 | by simp | 
| 0 | 119 | |
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
623diff
changeset | 120 | lemma nil_rsp[quot_respect]: | 
| 529 | 121 | assumes q: "Quotient R Abs Rep" | 
| 537 
57073b0b8fac
Even more name changes and cleaning
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
533diff
changeset | 122 | shows "list_rel R [] []" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 123 | by simp | 
| 0 | 124 | |
| 645 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 125 | lemma map_prs_aux: | 
| 540 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 126 | assumes a: "Quotient R1 abs1 rep1" | 
| 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 127 | and b: "Quotient R2 abs2 rep2" | 
| 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 128 | shows "(map abs2) (map ((abs1 ---> rep2) f) (map rep1 l)) = map f l" | 
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 129 | by (induct l) | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 130 | (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) | 
| 0 | 131 | |
| 645 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 132 | |
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 133 | lemma map_prs[quot_preserve]: | 
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 134 | assumes a: "Quotient R1 abs1 rep1" | 
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 135 | and b: "Quotient R2 abs2 rep2" | 
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 136 | shows "((abs1 ---> rep2) ---> (map rep1) ---> (map abs2)) map = map" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 137 | by (simp only: expand_fun_eq fun_map_def map_prs_aux[OF a b]) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 138 | (simp) | 
| 645 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 139 | |
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 140 | |
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
623diff
changeset | 141 | lemma map_rsp[quot_respect]: | 
| 529 | 142 | assumes q1: "Quotient R1 Abs1 Rep1" | 
| 143 | and q2: "Quotient R2 Abs2 Rep2" | |
| 565 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
541diff
changeset | 144 | shows "((R1 ===> R2) ===> (list_rel R1) ===> list_rel R2) map map" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 145 | apply(simp) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 146 | apply(rule allI)+ | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 147 | apply(rule impI) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 148 | apply(rule allI)+ | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 149 | apply (induct_tac xa ya rule: list_induct2') | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 150 | apply simp_all | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 151 | done | 
| 565 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
541diff
changeset | 152 | |
| 645 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 153 | lemma foldr_prs_aux: | 
| 540 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 154 | assumes a: "Quotient R1 abs1 rep1" | 
| 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 155 | and b: "Quotient R2 abs2 rep2" | 
| 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 156 | shows "abs2 (foldr ((abs1 ---> abs2 ---> rep2) f) (map rep1 l) (rep2 e)) = foldr f l e" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 157 | by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) | 
| 540 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 158 | |
| 796 
64f9c76f70c7
corrected wrong [quot_respect] attribute; tuned
 Christian Urban <urbanc@in.tum.de> parents: 
768diff
changeset | 159 | lemma foldr_prs[quot_preserve]: | 
| 645 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 160 | assumes a: "Quotient R1 abs1 rep1" | 
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 161 | and b: "Quotient R2 abs2 rep2" | 
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 162 | shows "((abs1 ---> abs2 ---> rep2) ---> (map rep1) ---> rep2 ---> abs2) foldr = foldr" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 163 | by (simp only: expand_fun_eq fun_map_def foldr_prs_aux[OF a b]) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 164 | (simp) | 
| 645 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 165 | |
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 166 | lemma foldl_prs_aux: | 
| 540 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 167 | assumes a: "Quotient R1 abs1 rep1" | 
| 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 168 | and b: "Quotient R2 abs2 rep2" | 
| 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 169 | shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 170 | by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b]) | 
| 645 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 171 | |
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 172 | |
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 173 | lemma foldl_prs[quot_preserve]: | 
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 174 | assumes a: "Quotient R1 abs1 rep1" | 
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 175 | and b: "Quotient R2 abs2 rep2" | 
| 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 176 | shows "((abs1 ---> abs2 ---> rep1) ---> rep1 ---> (map rep2) ---> abs1) foldl = foldl" | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 177 | by (simp only: expand_fun_eq fun_map_def foldl_prs_aux[OF a b]) | 
| 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 178 | (simp) | 
| 645 
fe2a37cfecd3
proper formulation of all preservation theorems
 Christian Urban <urbanc@in.tum.de> parents: 
644diff
changeset | 179 | |
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 180 | lemma list_rel_empty: | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 181 | shows "list_rel R [] b \<Longrightarrow> length b = 0" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 182 | by (induct b) (simp_all) | 
| 0 | 183 | |
| 1128 
17ca92ab4660
Main renaming + fixes for new Isabelle in IntEx2.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
937diff
changeset | 184 | lemma list_rel_len: | 
| 935 
c96e007b512f
cleaning of QuotProd; a little cleaning of QuotList
 Christian Urban <urbanc@in.tum.de> parents: 
924diff
changeset | 185 | shows "list_rel R a b \<Longrightarrow> length a = length b" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 186 | apply (induct a arbitrary: b) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 187 | apply (simp add: list_rel_empty) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 188 | apply (case_tac b) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 189 | apply simp_all | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 190 | done | 
| 0 | 191 | |
| 668 | 192 | (* induct_tac doesn't accept 'arbitrary', so we manually 'spec' *) | 
| 636 
520a4084d064
changed names of attributes
 Christian Urban <urbanc@in.tum.de> parents: 
623diff
changeset | 193 | lemma foldl_rsp[quot_respect]: | 
| 565 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
541diff
changeset | 194 | assumes q1: "Quotient R1 Abs1 Rep1" | 
| 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
541diff
changeset | 195 | and q2: "Quotient R2 Abs2 Rep2" | 
| 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
541diff
changeset | 196 | shows "((R1 ===> R2 ===> R1) ===> R1 ===> list_rel R2 ===> R1) foldl foldl" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 197 | apply(auto) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 198 | apply (subgoal_tac "R1 xa ya \<longrightarrow> list_rel R2 xb yb \<longrightarrow> R1 (foldl x xa xb) (foldl y ya yb)") | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 199 | apply simp | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 200 | apply (rule_tac x="xa" in spec) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 201 | apply (rule_tac x="ya" in spec) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 202 | apply (rule_tac xs="xb" and ys="yb" in list_induct2) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 203 | apply (rule list_rel_len) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 204 | apply (simp_all) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 205 | done | 
| 565 
baff284c6fcc
Proved foldl_rsp and ho_map_rsp
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
541diff
changeset | 206 | |
| 666 | 207 | lemma foldr_rsp[quot_respect]: | 
| 208 | assumes q1: "Quotient R1 Abs1 Rep1" | |
| 209 | and q2: "Quotient R2 Abs2 Rep2" | |
| 210 | shows "((R1 ===> R2 ===> R2) ===> list_rel R1 ===> R2 ===> R2) foldr foldr" | |
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 211 | apply auto | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 212 | apply(subgoal_tac "R2 xb yb \<longrightarrow> list_rel R1 xa ya \<longrightarrow> R2 (foldr x xa xb) (foldr y ya yb)") | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 213 | apply simp | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 214 | apply (rule_tac xs="xa" and ys="ya" in list_induct2) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 215 | apply (rule list_rel_len) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 216 | apply (simp_all) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 217 | done | 
| 540 
c0b13fb70d6d
More code cleaning and renaming: moved rsp and prs lemmas from Int to QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
539diff
changeset | 218 | |
| 825 
970e86082cd7
Modifictaions for new_relation.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
796diff
changeset | 219 | lemma list_rel_eq[id_simps]: | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 220 | shows "(list_rel (op =)) = (op =)" | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 221 | unfolding expand_fun_eq | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 222 | apply(rule allI)+ | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 223 | apply(induct_tac x xa rule: list_induct2') | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 224 | apply(simp_all) | 
| 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 225 | done | 
| 0 | 226 | |
| 541 
94deffed619d
more name cleaning and removing
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
540diff
changeset | 227 | lemma list_rel_refl: | 
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 228 | assumes a: "\<And>x y. R x y = (R x = R y)" | 
| 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 229 | shows "list_rel R x x" | 
| 927 
04ef4bd3b936
more eq_reflection & other cleaning.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
924diff
changeset | 230 | by (induct x) (auto simp add: a) | 
| 0 | 231 | |
| 539 
8287fb5b8d7a
Cleaning & Renaming coming from QuotList
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
537diff
changeset | 232 | end |